reworked Probability theory
authorhoelzl
Mon, 23 Apr 2012 12:14:35 +0200
changeset 47694 05663f75964c
parent 47693 64023cf4d148
child 47695 89a90da149a9
reworked Probability theory
NEWS
src/HOL/IsaMakefile
src/HOL/Probability/Binary_Product_Measure.thy
src/HOL/Probability/Borel_Space.thy
src/HOL/Probability/Caratheodory.thy
src/HOL/Probability/Complete_Measure.thy
src/HOL/Probability/Conditional_Probability.thy
src/HOL/Probability/Finite_Product_Measure.thy
src/HOL/Probability/Independent_Family.thy
src/HOL/Probability/Infinite_Product_Measure.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/Measure_Space.thy
src/HOL/Probability/Probability.thy
src/HOL/Probability/Probability_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/NEWS	Mon Apr 23 12:23:23 2012 +0100
+++ b/NEWS	Mon Apr 23 12:14:35 2012 +0200
@@ -647,6 +647,180 @@
 
   DERIV_nonneg_imp_nonincreasing ~> DERIV_nonneg_imp_nondecreasing
 
+* Theory Library/Multiset: Improved code generation of multisets.
+
+* Session HOL-Probability: Introduced the type "'a measure" to represent
+measures, this replaces the records 'a algebra and 'a measure_space. The
+locales based on subset_class now have two locale-parameters the space
+\<Omega> and the set of measurables sets M. The product of probability spaces
+uses now the same constant as the finite product of sigma-finite measure
+spaces "PiM :: ('i => 'a) measure". Most constants are defined now
+outside of locales and gain an additional parameter, like null_sets,
+almost_eventually or \<mu>'. Measure space constructions for distributions
+and densities now got their own constants distr and density. Instead of
+using locales to describe measure spaces with a finite space, the
+measure count_space and point_measure is introduced. INCOMPATIBILITY.
+
+  Renamed constants:
+  measure -> emeasure
+  finite_measure.\<mu>' -> measure
+  product_algebra_generator -> prod_algebra
+  product_prob_space.emb -> prod_emb
+  product_prob_space.infprod_algebra -> PiM
+
+  Removed locales:
+  completeable_measure_space
+  finite_measure_space
+  finite_prob_space
+  finite_product_finite_prob_space
+  finite_product_sigma_algebra
+  finite_sigma_algebra
+  measure_space
+  pair_finite_prob_space
+  pair_finite_sigma_algebra
+  pair_finite_space
+  pair_sigma_algebra
+  product_sigma_algebra
+
+  Removed constants:
+  distribution -> use distr measure, or distributed predicate
+  joint_distribution -> use distr measure, or distributed predicate
+  product_prob_space.infprod_algebra -> use PiM
+  subvimage
+  image_space
+  conditional_space
+  pair_measure_generator
+
+  Replacement theorems:
+  sigma_algebra.measurable_sigma -> measurable_measure_of
+  measure_space.additive -> emeasure_additive
+  measure_space.measure_additive -> plus_emeasure
+  measure_space.measure_mono -> emeasure_mono
+  measure_space.measure_top -> emeasure_space
+  measure_space.measure_compl -> emeasure_compl
+  measure_space.measure_Diff -> emeasure_Diff
+  measure_space.measure_countable_increasing -> emeasure_countable_increasing
+  measure_space.continuity_from_below -> SUP_emeasure_incseq
+  measure_space.measure_incseq -> incseq_emeasure
+  measure_space.continuity_from_below_Lim -> Lim_emeasure_incseq
+  measure_space.measure_decseq -> decseq_emeasure
+  measure_space.continuity_from_above -> INF_emeasure_decseq
+  measure_space.measure_insert -> emeasure_insert
+  measure_space.measure_setsum -> setsum_emeasure
+  measure_space.measure_finite_singleton -> emeasure_eq_setsum_singleton
+  finite_additivity_sufficient -> ring_of_sets.countably_additiveI_finite
+  measure_space.measure_setsum_split -> setsum_emeasure_cover
+  measure_space.measure_subadditive -> subadditive
+  measure_space.measure_subadditive_finite -> emeasure_subadditive_finite
+  measure_space.measure_eq_0 -> emeasure_eq_0
+  measure_space.measure_finitely_subadditive -> emeasure_subadditive_finite
+  measure_space.measure_countably_subadditive -> emeasure_subadditive_countably
+  measure_space.measure_UN_eq_0 -> emeasure_UN_eq_0
+  measure_unique_Int_stable -> measure_eqI_generator_eq
+  measure_space.measure_Diff_null_set -> emeasure_Diff_null_set
+  measure_space.measure_Un_null_set -> emeasure_Un_null_set
+  measure_space.almost_everywhere_def -> eventually_ae_filter
+  measure_space.almost_everywhere_vimage -> AE_distrD
+  measure_space.measure_space_vimage -> emeasure_distr
+  measure_space.AE_iff_null_set -> AE_iff_null
+  measure_space.real_measure_Union -> measure_Union
+  measure_space.real_measure_finite_Union -> measure_finite_Union
+  measure_space.real_measure_Diff -> measure_Diff
+  measure_space.real_measure_UNION -> measure_UNION
+  measure_space.real_measure_subadditive -> measure_subadditive
+  measure_space.real_measure_setsum_singleton -> measure_eq_setsum_singleton
+  measure_space.real_continuity_from_below -> Lim_measure_incseq
+  measure_space.continuity_from_above_Lim -> Lim_emeasure_decseq
+  measure_space.real_continuity_from_above -> Lim_measure_decseq
+  measure_space.real_measure_countably_subadditive -> measure_subadditive_countably
+  finite_measure.finite_measure -> finite_measure.emeasure_finite
+  finite_measure.finite_measure_eq -> finite_measure.emeasure_eq_measure
+  finite_measure.positive_measure' -> measure_nonneg
+  finite_measure.real_measure -> finite_measure.emeasure_real
+  finite_measure.empty_measure -> measure_empty
+  finite_measure.finite_measure_countably_subadditive -> finite_measure.finite_measure_subadditive_countably
+  finite_measure.finite_measure_finite_singleton -> finite_measure.finite_measure_eq_setsum_singleton
+  finite_measure.finite_continuity_from_below -> finite_measure.finite_Lim_measure_incseq
+  finite_measure.finite_continuity_from_above -> finite_measure.finite_Lim_measure_decseq
+  measure_space.simple_integral_vimage -> simple_integral_distr
+  measure_space.integrable_vimage -> integrable_distr
+  measure_space.positive_integral_translated_density -> positive_integral_density
+  measure_space.integral_translated_density -> integral_density
+  measure_space.integral_vimage -> integral_distr
+  measure_space_density -> emeasure_density
+  measure_space.positive_integral_vimage -> positive_integral_distr
+  measure_space.simple_function_vimage -> simple_function_comp
+  measure_space.simple_integral_vimage -> simple_integral_distr
+  pair_sigma_algebra.measurable_cut_fst -> sets_Pair1
+  pair_sigma_algebra.measurable_cut_snd -> sets_Pair2
+  pair_sigma_algebra.measurable_pair_image_fst -> measurable_Pair1
+  pair_sigma_algebra.measurable_pair_image_snd -> measurable_Pair2
+  pair_sigma_algebra.measurable_product_swap -> measurable_pair_swap_iff
+  pair_sigma_finite.measure_cut_measurable_fst -> pair_sigma_finite.measurable_emeasure_Pair1
+  pair_sigma_finite.measure_cut_measurable_snd -> pair_sigma_finite.measurable_emeasure_Pair2
+  measure_space.measure_not_negative -> emeasure_not_MInf
+  pair_sigma_finite.measure_preserving_swap -> pair_sigma_finite.distr_pair_swap
+  pair_sigma_finite.pair_measure_alt -> pair_sigma_finite.emeasure_pair_measure_alt
+  pair_sigma_finite.pair_measure_alt2 -> pair_sigma_finite.emeasure_pair_measure_alt2
+  pair_sigma_finite.pair_measure_times -> pair_sigma_finite.emeasure_pair_measure_Times
+  pair_sigma_algebra.pair_sigma_algebra_measurable -> measurable_pair_swap
+  pair_sigma_algebra.pair_sigma_algebra_swap_measurable -> measurable_pair_swap'
+  pair_sigma_algebra.sets_swap -> sets_pair_swap
+  finite_product_sigma_algebra.in_P -> sets_PiM_I_finite
+  Int_stable_product_algebra_generator -> positive_integral
+  product_sigma_finite.measure_fold -> product_sigma_finite.distr_merge
+  product_sigma_finite.measure_preserving_component_singelton -> product_sigma_finite.distr_singleton
+  product_sigma_finite.measure_preserving_merge -> product_sigma_finite.distr_merge
+  finite_product_sigma_algebra.P_empty -> space_PiM_empty, sets_PiM_empty
+  product_algebra_generator_der -> prod_algebra_eq_finite
+  product_algebra_generator_into_space -> prod_algebra_sets_into_space
+  product_sigma_algebra.product_algebra_into_space -> space_closed
+  product_algebraE -> prod_algebraE_all
+  product_algebraI -> sets_PiM_I_finite
+  product_measure_exists -> product_sigma_finite.sigma_finite
+  sets_product_algebra -> sets_PiM
+  sigma_product_algebra_sigma_eq -> sigma_prod_algebra_sigma_eq
+  space_product_algebra -> space_PiM
+  Int_stable_cuboids -> Int_stable_atLeastAtMost
+  measure_space.density_is_absolutely_continuous -> absolutely_continuousI_density
+  sigma_finite_measure.RN_deriv_vimage -> sigma_finite_measure.RN_deriv_distr
+  prob_space_unique_Int_stable -> measure_eqI_prob_space
+  sigma_finite_measure.disjoint_sigma_finite -> sigma_finite_disjoint
+  prob_space.measure_space_1 -> prob_space.emeasure_space_1
+  prob_space.prob_space_vimage -> prob_space_distr
+  prob_space.random_variable_restrict -> measurable_restrict
+  measure_preserving -> equality "distr M N f = N" "f : measurable M N"
+  measure_unique_Int_stable_vimage -> measure_eqI_generator_eq
+  measure_space.measure_preserving_Int_stable -> measure_eqI_generator_eq
+  product_prob_space.finite_index_eq_finite_product -> product_prob_space.sets_PiM_generator
+  product_prob_space.finite_measure_infprod_emb_Pi -> product_prob_space.measure_PiM_emb
+  finite_product_prob_space.finite_measure_times -> finite_product_prob_space.finite_measure_PiM_emb
+  product_prob_space.infprod_spec -> product_prob_space.emeasure_PiM_emb_not_empty
+  product_prob_space.measurable_component -> measurable_component_singleton
+  product_prob_space.measurable_emb -> measurable_prod_emb
+  product_prob_space.measurable_into_infprod_algebra -> measurable_PiM_single
+  product_prob_space.measurable_singleton_infprod -> measurable_component_singleton
+  product_prob_space.measure_emb -> emeasure_prod_emb
+  sequence_space.measure_infprod -> sequence_space.measure_PiM_countable
+  product_prob_space.measure_preserving_restrict -> product_prob_space.distr_restrict
+  prob_space.indep_distribution_eq_measure -> prob_space.indep_vars_iff_distr_eq_PiM
+  prob_space.indep_var_distributionD -> prob_space.indep_var_distribution_eq
+  conditional_entropy_positive -> conditional_entropy_nonneg_simple
+  conditional_entropy_eq -> conditional_entropy_simple_distributed
+  conditional_mutual_information_eq_mutual_information -> conditional_mutual_information_eq_mutual_information_simple
+  conditional_mutual_information_generic_positive -> conditional_mutual_information_nonneg_simple
+  conditional_mutual_information_positive -> conditional_mutual_information_nonneg_simple
+  entropy_commute -> entropy_commute_simple
+  entropy_eq -> entropy_simple_distributed
+  entropy_generic_eq -> entropy_simple_distributed
+  entropy_positive -> entropy_nonneg_simple
+  entropy_uniform_max -> entropy_uniform
+  KL_eq_0 -> KL_same_eq_0
+  KL_eq_0_imp -> KL_eq_0_iff_eq
+  KL_ge_0 -> KL_nonneg
+  mutual_information_eq -> mutual_information_simple_distributed
+  mutual_information_positive -> mutual_information_nonneg_simple
+
 * New "case_product" attribute to generate a case rule doing multiple
 case distinctions at the same time.  E.g.
 
@@ -655,8 +829,6 @@
 produces a rule which can be used to perform case distinction on both
 a list and a nat.
 
-* Theory Library/Multiset: Improved code generation of multisets.
-
 * New Transfer package:
 
   - transfer_rule attribute: Maintains a collection of transfer rules,
--- a/src/HOL/IsaMakefile	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/IsaMakefile	Mon Apr 23 12:14:35 2012 +0200
@@ -1198,14 +1198,13 @@
 $(OUT)/HOL-Probability: $(OUT)/HOL-Multivariate_Analysis		\
   Probability/Binary_Product_Measure.thy Probability/Borel_Space.thy	\
   Probability/Caratheodory.thy Probability/Complete_Measure.thy		\
-  Probability/Conditional_Probability.thy				\
   Probability/ex/Dining_Cryptographers.thy				\
   Probability/ex/Koepf_Duermuth_Countermeasure.thy			\
   Probability/Finite_Product_Measure.thy				\
   Probability/Independent_Family.thy					\
   Probability/Infinite_Product_Measure.thy Probability/Information.thy	\
   Probability/Lebesgue_Integration.thy Probability/Lebesgue_Measure.thy \
-  Probability/Measure.thy Probability/Probability_Measure.thy		\
+  Probability/Measure_Space.thy Probability/Probability_Measure.thy	\
   Probability/Probability.thy Probability/Radon_Nikodym.thy		\
   Probability/ROOT.ML Probability/Sigma_Algebra.thy			\
   Library/Countable.thy Library/FuncSet.thy Library/Nat_Bijection.thy
--- a/src/HOL/Probability/Binary_Product_Measure.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Binary_Product_Measure.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -28,413 +28,317 @@
 
 section "Binary products"
 
-definition
-  "pair_measure_generator A B =
-    \<lparr> space = space A \<times> space B,
-      sets = {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B},
-      measure = \<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+y. indicator X (x,y) \<partial>B) \<partial>A \<rparr>"
-
 definition pair_measure (infixr "\<Otimes>\<^isub>M" 80) where
-  "A \<Otimes>\<^isub>M B = sigma (pair_measure_generator A B)"
-
-locale pair_sigma_algebra = M1: sigma_algebra M1 + M2: sigma_algebra M2
-  for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
-
-abbreviation (in pair_sigma_algebra)
-  "E \<equiv> pair_measure_generator M1 M2"
-
-abbreviation (in pair_sigma_algebra)
-  "P \<equiv> M1 \<Otimes>\<^isub>M M2"
-
-lemma sigma_algebra_pair_measure:
-  "sets M1 \<subseteq> Pow (space M1) \<Longrightarrow> sets M2 \<subseteq> Pow (space M2) \<Longrightarrow> sigma_algebra (pair_measure M1 M2)"
-  by (force simp: pair_measure_def pair_measure_generator_def intro!: sigma_algebra_sigma)
-
-sublocale pair_sigma_algebra \<subseteq> sigma_algebra P
-  using M1.space_closed M2.space_closed
-  by (rule sigma_algebra_pair_measure)
-
-lemma pair_measure_generatorI[intro, simp]:
-  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (pair_measure_generator A B)"
-  by (auto simp add: pair_measure_generator_def)
-
-lemma pair_measureI[intro, simp]:
-  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)"
-  by (auto simp add: pair_measure_def)
+  "A \<Otimes>\<^isub>M B = measure_of (space A \<times> space B)
+      {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
+      (\<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+y. indicator X (x,y) \<partial>B) \<partial>A)"
 
 lemma space_pair_measure:
   "space (A \<Otimes>\<^isub>M B) = space A \<times> space B"
-  by (simp add: pair_measure_def pair_measure_generator_def)
+  unfolding pair_measure_def using space_closed[of A] space_closed[of B]
+  by (intro space_measure_of) auto
+
+lemma sets_pair_measure:
+  "sets (A \<Otimes>\<^isub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
+  unfolding pair_measure_def using space_closed[of A] space_closed[of B]
+  by (intro sets_measure_of) auto
 
-lemma sets_pair_measure_generator:
-  "sets (pair_measure_generator N M) = (\<lambda>(x, y). x \<times> y) ` (sets N \<times> sets M)"
-  unfolding pair_measure_generator_def by auto
+lemma pair_measureI[intro, simp]:
+  "x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)"
+  by (auto simp: sets_pair_measure)
 
-lemma pair_measure_generator_sets_into_space:
-  assumes "sets M \<subseteq> Pow (space M)" "sets N \<subseteq> Pow (space N)"
-  shows "sets (pair_measure_generator M N) \<subseteq> Pow (space (pair_measure_generator M N))"
-  using assms by (auto simp: pair_measure_generator_def)
+lemma measurable_pair_measureI:
+  assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2"
+  assumes 2: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> f -` (A \<times> B) \<inter> space M \<in> sets M"
+  shows "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
+  unfolding pair_measure_def using 1 2
+  by (intro measurable_measure_of) (auto dest: sets_into_space)
 
-lemma pair_measure_generator_Int_snd:
-  assumes "sets S1 \<subseteq> Pow (space S1)"
-  shows "sets (pair_measure_generator S1 (algebra.restricted_space S2 A)) =
-         sets (algebra.restricted_space (pair_measure_generator S1 S2) (space S1 \<times> A))"
-  (is "?L = ?R")
-  apply (auto simp: pair_measure_generator_def image_iff)
-  using assms
-  apply (rule_tac x="a \<times> xa" in exI)
-  apply force
-  using assms
-  apply (rule_tac x="a" in exI)
-  apply (rule_tac x="b \<inter> A" in exI)
-  apply auto
-  done
+lemma measurable_fst[intro!, simp]: "fst \<in> measurable (M1 \<Otimes>\<^isub>M M2) M1"
+  unfolding measurable_def
+proof safe
+  fix A assume A: "A \<in> sets M1"
+  from this[THEN sets_into_space] have "fst -` A \<inter> space M1 \<times> space M2 = A \<times> space M2" by auto
+  with A show "fst -` A \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by (simp add: space_pair_measure)
+qed (simp add: space_pair_measure)
 
-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 -
-  { 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"])
-      using M1.sets_into_space by force+ }
-  moreover
-  { fix X assume "X \<in> sets M2"
-    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 have "?fst \<and> ?snd"
-    by (fastforce simp: measurable_def sets_sigma space_pair_measure
-                 intro!: sigma_sets.Basic)
-  then show ?fst ?snd by auto
-qed
+lemma measurable_snd[intro!, simp]: "snd \<in> measurable (M1 \<Otimes>\<^isub>M M2) M2"
+  unfolding measurable_def
+proof safe
+  fix A assume A: "A \<in> sets M2"
+  from this[THEN sets_into_space] have "snd -` A \<inter> space M1 \<times> space M2 = space M1 \<times> A" by auto
+  with A show "snd -` A \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by (simp add: space_pair_measure)
+qed (simp add: space_pair_measure)
+
+lemma measurable_fst': "f \<in> measurable M (N \<Otimes>\<^isub>M P) \<Longrightarrow> (\<lambda>x. fst (f x)) \<in> measurable M N"
+  using measurable_comp[OF _ measurable_fst] by (auto simp: comp_def)
+
+lemma measurable_snd': "f \<in> measurable M (N \<Otimes>\<^isub>M P) \<Longrightarrow> (\<lambda>x. snd (f x)) \<in> measurable M P"
+    using measurable_comp[OF _ measurable_snd] by (auto simp: comp_def)
 
-lemma (in pair_sigma_algebra) measurable_pair_iff:
-  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" unfolding pair_measure_def
-    proof (rule M.measurable_sigma)
-      show "sets (pair_measure_generator M1 M2) \<subseteq> Pow (space E)"
-        unfolding pair_measure_generator_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 pair_measure_generator_def)
-      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_measure_generator_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
+lemma measurable_fst'': "f \<in> measurable M N \<Longrightarrow> (\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^isub>M P) N"
+  using measurable_comp[OF measurable_fst _] by (auto simp: comp_def)
+
+lemma measurable_snd'': "f \<in> measurable M N \<Longrightarrow> (\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^isub>M M) N"
+  using measurable_comp[OF measurable_snd _] by (auto simp: comp_def)
+
+lemma measurable_pair_iff:
+  "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
+proof safe
+  assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
+  show "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
+  proof (rule measurable_pair_measureI)
+    show "f \<in> space M \<rightarrow> space M1 \<times> space M2"
+      using f s by (auto simp: mem_Times_iff measurable_def comp_def)
+    fix A B assume "A \<in> sets M1" "B \<in> sets M2"
+    moreover have "(fst \<circ> f) -` A \<inter> space M \<in> sets M" "(snd \<circ> f) -` B \<inter> space M \<in> sets M"
+      using f `A \<in> sets M1` s `B \<in> sets M2` by (auto simp: measurable_sets)
+    moreover have "f -` (A \<times> B) \<inter> space M = ((fst \<circ> f) -` A \<inter> space M) \<inter> ((snd \<circ> f) -` B \<inter> space M)"
+      by (auto simp: vimage_Times)
+    ultimately show "f -` (A \<times> B) \<inter> space M \<in> sets M" by auto
   qed
-qed
+qed auto
 
-lemma (in pair_sigma_algebra) measurable_pair:
-  assumes "sigma_algebra M"
-  assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"
-  shows "f \<in> measurable M P"
-  unfolding measurable_pair_iff[OF assms(1)] using assms(2,3) by simp
-
-lemma pair_measure_generatorE:
-  assumes "X \<in> sets (pair_measure_generator M1 M2)"
-  obtains A B where "X = A \<times> B" "A \<in> sets M1" "B \<in> sets M2"
-  using assms unfolding pair_measure_generator_def by auto
+lemma measurable_pair:
+  "(fst \<circ> f) \<in> measurable M M1 \<Longrightarrow> (snd \<circ> f) \<in> measurable M M2 \<Longrightarrow> f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)"
+  unfolding measurable_pair_iff by simp
 
-lemma (in pair_sigma_algebra) pair_measure_generator_swap:
-  "(\<lambda>X. (\<lambda>(x,y). (y,x)) -` X \<inter> space M2 \<times> space M1) ` sets E = sets (pair_measure_generator M2 M1)"
-proof (safe elim!: pair_measure_generatorE)
-  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_measure_generator M2 M1)"
-    by (auto intro: pair_measure_generatorI)
-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_measure_generatorI)
+lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
+proof (rule measurable_pair_measureI)
+  fix A B assume "A \<in> sets M2" "B \<in> sets M1"
+  moreover then have "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space (M1 \<Otimes>\<^isub>M M2) = B \<times> A"
+    by (auto dest: sets_into_space simp: space_pair_measure)
+  ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+    by auto
+qed (auto simp add: space_pair_measure)
+
+lemma measurable_pair_swap:
+  assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^isub>M M1) M"
+proof -
+  have "(\<lambda>x. f (case x of (x, y) \<Rightarrow> (y, x))) = (\<lambda>(x, y). f (y, x))" by auto
+  then show ?thesis
+    using measurable_comp[OF measurable_pair_swap' f] by (simp add: comp_def)
 qed
 
-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 (M2 \<Otimes>\<^isub>M M1)" (is "_ \<in> sets ?Q")
-proof -
-  let ?f = "\<lambda>Q. (\<lambda>(x,y). (y, x)) -` Q \<inter> space M2 \<times> space M1"
-  have *: "(\<lambda>(x,y). (y, x)) -` Q = ?f Q"
-    using sets_into_space[OF Q] by (auto simp: space_pair_measure)
-  have "sets (M2 \<Otimes>\<^isub>M M1) = sets (sigma (pair_measure_generator M2 M1))"
-    unfolding pair_measure_def ..
-  also have "\<dots> = sigma_sets (space M2 \<times> space M1) (?f ` sets E)"
-    unfolding sigma_def pair_measure_generator_swap[symmetric]
-    by (simp add: pair_measure_generator_def)
-  also have "\<dots> = ?f ` sigma_sets (space M1 \<times> space M2) (sets E)"
-    using M1.sets_into_space M2.sets_into_space
-    by (intro sigma_sets_vimage) (auto simp: pair_measure_generator_def)
-  also have "\<dots> = ?f ` sets P"
-    unfolding pair_measure_def pair_measure_generator_def sigma_def by simp
-  finally show ?thesis
-    using Q by (subst *) auto
-qed
+lemma measurable_pair_swap_iff:
+  "f \<in> measurable (M2 \<Otimes>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) M"
+  using measurable_pair_swap[of "\<lambda>(x,y). f (y, x)"]
+  by (auto intro!: measurable_pair_swap)
 
-lemma (in pair_sigma_algebra) pair_sigma_algebra_swap_measurable:
-  shows "(\<lambda>(x,y). (y, x)) \<in> measurable P (M2 \<Otimes>\<^isub>M 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_measure_generator_def pair_measure_def by auto
-next
-  fix A assume "A \<in> sets (M2 \<Otimes>\<^isub>M M1)"
-  interpret Q: pair_sigma_algebra M2 M1 by default
-  from Q.sets_pair_sigma_algebra_swap[OF `A \<in> sets (M2 \<Otimes>\<^isub>M M1)`]
-  show "?f -` A \<inter> space ?P \<in> sets ?P" by simp
-qed
+lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^isub>M M2)"
+proof (rule measurable_pair_measureI)
+  fix A B assume "A \<in> sets M1" "B \<in> sets M2"
+  moreover then have "Pair x -` (A \<times> B) \<inter> space M2 = (if x \<in> A then B else {})"
+    by (auto dest: sets_into_space simp: space_pair_measure)
+  ultimately show "Pair x -` (A \<times> B) \<inter> space M2 \<in> sets M2"
+    by auto
+qed (auto simp add: space_pair_measure)
 
-lemma (in pair_sigma_algebra) measurable_cut_fst[simp,intro]:
-  assumes "Q \<in> sets P" shows "Pair x -` Q \<in> sets M2"
+lemma sets_Pair1: assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "Pair x -` A \<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_measure)
-  have "sets E \<subseteq> sets ?Q"
-    using M1.sets_into_space M2.sets_into_space
-    by (auto simp: pair_measure_generator_def space_pair_measure)
-  then have "sets P \<subseteq> sets ?Q"
-    apply (subst pair_measure_def, intro Q.sets_sigma_subset)
-    by (simp add: pair_measure_def)
-  with assms show ?thesis by auto
+  have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})"
+    using A[THEN sets_into_space] by (auto simp: space_pair_measure)
+  also have "\<dots> \<in> sets M2"
+    using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: split_if_asm)
+  finally show ?thesis .
 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
-  from Q.measurable_cut_fst[OF sets_pair_sigma_algebra_swap[OF Q], of y]
-  show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
-qed
+lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^isub>M M2)"
+  using measurable_comp[OF measurable_Pair1' measurable_pair_swap', of y M2 M1]
+  by (simp add: comp_def)
 
-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_measure_generator_def pair_measure_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_measure_generator_def pair_measure_def)
-  finally show "(\<lambda>y. f (x, y)) -` A \<inter> space M2 \<in> sets M2" .
+lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"
+proof -
+  have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})"
+    using A[THEN sets_into_space] by (auto simp: space_pair_measure)
+  also have "\<dots> \<in> sets M1"
+    using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: split_if_asm)
+  finally show ?thesis .
 qed
 
-lemma (in pair_sigma_algebra) measurable_pair_image_fst:
-  assumes m: "f \<in> measurable P M" and "y \<in> space M2"
+lemma measurable_Pair2:
+  assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" and x: "x \<in> space M1"
+  shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"
+  using measurable_comp[OF measurable_Pair1' f, OF x]
+  by (simp add: comp_def)
+  
+lemma measurable_Pair1:
+  assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" and y: "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
+  using measurable_comp[OF measurable_Pair2' f, OF y]
+  by (simp add: comp_def)
 
-lemma (in pair_sigma_algebra) Int_stable_pair_measure_generator: "Int_stable E"
+lemma Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
   unfolding Int_stable_def
-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_measure_generator_def by auto
-  then show "A \<inter> B \<in> sets E"
-    by (auto simp add: times_Int_times pair_measure_generator_def)
-qed
+  by safe (auto simp add: times_Int_times)
 
 lemma finite_measure_cut_measurable:
-  fixes M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
   assumes "sigma_finite_measure M1" "finite_measure M2"
   assumes "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)"
-  shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1"
+  shows "(\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"
     (is "?s Q \<in> _")
 proof -
   interpret M1: sigma_finite_measure M1 by fact
   interpret M2: finite_measure M2 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>"
+  let ?\<Omega> = "space M1 \<times> space M2" and ?D = "{A\<in>sets (M1 \<Otimes>\<^isub>M M2). ?s A \<in> borel_measurable M1}"
   note space_pair_measure[simp]
-  interpret dynkin_system ?D
+  interpret dynkin_system ?\<Omega> ?D
   proof (intro dynkin_systemI)
-    fix A assume "A \<in> sets ?D" then show "A \<subseteq> space ?D"
-      using sets_into_space by simp
+    fix A assume "A \<in> ?D" then show "A \<subseteq> ?\<Omega>"
+      using sets_into_space[of A "M1 \<Otimes>\<^isub>M M2"] by simp
   next
-    from top show "space ?D \<in> sets ?D"
-      by (auto simp add: if_distrib intro!: M1.measurable_If)
+    from top show "?\<Omega> \<in> ?D"
+      by (auto simp add: if_distrib intro!: measurable_If)
   next
-    fix A assume "A \<in> sets ?D"
-    with sets_into_space have "\<And>x. measure M2 (Pair x -` (space M1 \<times> space M2 - A)) =
-        (if x \<in> space M1 then measure M2 (space M2) - ?s A x else 0)"
-      by (auto intro!: M2.measure_compl 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_ereal_diff)
+    fix A assume "A \<in> ?D"
+    with sets_into_space have "\<And>x. emeasure M2 (Pair x -` (?\<Omega> - A)) =
+        (if x \<in> space M1 then emeasure M2 (space M2) - ?s A x else 0)"
+      by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1)
+    with `A \<in> ?D` top show "?\<Omega> - A \<in> ?D"
+      by (auto intro!: measurable_If)
   next
-    fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> sets ?D"
-    moreover then have "\<And>x. measure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
-      by (intro M2.measure_countably_additive[symmetric])
-         (auto simp: disjoint_family_on_def)
-    ultimately show "(\<Union>i. F i) \<in> sets ?D"
-      by (auto simp: vimage_UN intro!: M1.borel_measurable_psuminf)
+    fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> ?D"
+    moreover then have "\<And>x. emeasure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)"
+      by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def sets_Pair1)
+    ultimately show "(\<Union>i. F i) \<in> ?D"
+      by (auto simp: vimage_UN intro!: borel_measurable_psuminf)
   qed
-  have "sets P = sets ?D" apply (subst pair_measure_def)
-  proof (intro dynkin_lemma)
-    show "Int_stable E" by (rule Int_stable_pair_measure_generator)
-    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_measure_generator_def sets_sigma if_distrib
-               intro: sigma_sets.Basic intro!: M1.measurable_If)
-  qed (auto simp: pair_measure_def)
-  with `Q \<in> sets P` have "Q \<in> sets ?D" by simp
+  let ?G = "{a \<times> b | a b. a \<in> sets M1 \<and> b \<in> sets M2}"
+  have "sigma_sets ?\<Omega> ?G = ?D"
+  proof (rule dynkin_lemma)
+    show "?G \<subseteq> ?D"
+      by (auto simp: if_distrib Int_def[symmetric] intro!: measurable_If)
+  qed (auto simp: sets_pair_measure  Int_stable_pair_measure_generator)
+  with `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` have "Q \<in> ?D"
+    by (simp add: sets_pair_measure[symmetric])
   then show "?s Q \<in> borel_measurable M1" by simp
 qed
 
-subsection {* Binary products of $\sigma$-finite measure spaces *}
+subsection {* Binary products of $\sigma$-finite emeasure spaces *}
 
-locale pair_sigma_finite = pair_sigma_algebra M1 M2 + M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2
-  for M1 :: "('a, 'c) measure_space_scheme" and M2 :: "('b, 'd) measure_space_scheme"
+locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2
+  for M1 :: "'a measure" and M2 :: "'b measure"
 
-lemma (in pair_sigma_finite) measure_cut_measurable_fst:
-  assumes "Q \<in> sets P" shows "(\<lambda>x. measure M2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
+lemma sets_pair_measure_cong[cong]:
+  "sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^isub>M M2) = sets (M1' \<Otimes>\<^isub>M M2')"
+  unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)
+
+lemma (in pair_sigma_finite) measurable_emeasure_Pair1:
+  assumes Q: "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _")
 proof -
-  have [intro]: "sigma_algebra M1" and [intro]: "sigma_algebra M2" by default+
-  have M1: "sigma_finite_measure M1" by default
-  from M2.disjoint_sigma_finite guess F .. note F = this
+  from M2.sigma_finite_disjoint guess F . note F = this
   then have F_sets: "\<And>i. F i \<in> sets M2" by auto
+  have M1: "sigma_finite_measure M1" ..
   let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q"
   { fix i
-    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
-    let ?R2 = "M2.restricted_space (F i)"
-    have "(\<lambda>x. measure ?R2 (Pair x -` (space M1 \<times> space ?R2 \<inter> Q))) \<in> borel_measurable M1"
+      using F sets_into_space by auto
+    let ?R = "density M2 (indicator (F i))"
+    have "(\<lambda>x. emeasure ?R (Pair x -` (space M1 \<times> space ?R \<inter> Q))) \<in> borel_measurable M1"
     proof (intro finite_measure_cut_measurable[OF M1])
-      show "finite_measure ?R2"
-        using F by (intro M2.restricted_to_finite_measure) auto
-      have "(space M1 \<times> space ?R2) \<inter> Q \<in> (op \<inter> (space M1 \<times> F i)) ` sets P"
-        using `Q \<in> sets P` by (auto simp: image_iff)
-      also have "\<dots> = sigma_sets (space M1 \<times> F i) ((op \<inter> (space M1 \<times> F i)) ` sets E)"
-        unfolding pair_measure_def pair_measure_generator_def sigma_def
-        using `F i \<in> sets M2` M2.sets_into_space
-        by (auto intro!: sigma_sets_Int sigma_sets.Basic)
-      also have "\<dots> \<subseteq> sets (M1 \<Otimes>\<^isub>M ?R2)"
-        using M1.sets_into_space
-        apply (auto simp: times_Int_times pair_measure_def pair_measure_generator_def sigma_def
-                    intro!: sigma_sets_subseteq)
-        apply (rule_tac x="a" in exI)
-        apply (rule_tac x="b \<inter> F i" in exI)
-        by auto
-      finally show "(space M1 \<times> space ?R2) \<inter> Q \<in> sets (M1 \<Otimes>\<^isub>M ?R2)" .
+      show "finite_measure ?R"
+        using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq)
+      show "(space M1 \<times> space ?R) \<inter> Q \<in> sets (M1 \<Otimes>\<^isub>M ?R)"
+        using Q by (simp add: Int)
     qed
-    moreover have "\<And>x. 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_measure)
-    ultimately have "(\<lambda>x. measure M2 (?C x i)) \<in> borel_measurable M1"
+    moreover have "\<And>x. emeasure ?R (Pair x -` (space M1 \<times> space ?R \<inter> Q))
+        = emeasure M2 (F i \<inter> Pair x -` (space M1 \<times> space ?R \<inter> Q))"
+      using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1)
+    moreover have "\<And>x. F i \<inter> Pair x -` (space M1 \<times> space ?R \<inter> Q) = ?C x i"
+      using sets_into_space[OF Q] by (auto simp: space_pair_measure)
+    ultimately have "(\<lambda>x. emeasure M2 (?C x i)) \<in> borel_measurable M1"
       by simp }
   moreover
   { fix x
-    have "(\<Sum>i. measure M2 (?C x i)) = measure M2 (\<Union>i. ?C x i)"
-    proof (intro M2.measure_countably_additive)
+    have "(\<Sum>i. emeasure M2 (?C x i)) = emeasure M2 (\<Union>i. ?C x i)"
+    proof (intro suminf_emeasure)
       show "range (?C x) \<subseteq> sets M2"
-        using F `Q \<in> sets P` by (auto intro!: M2.Int)
+        using F `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` by (auto intro!: sets_Pair1)
       have "disjoint_family F" using F by auto
       show "disjoint_family (?C x)"
         by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto
     qed
     also have "(\<Union>i. ?C x i) = Pair x -` Q"
-      using F sets_into_space `Q \<in> sets P`
+      using F sets_into_space[OF `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)`]
       by (auto simp: space_pair_measure)
-    finally have "measure M2 (Pair x -` Q) = (\<Sum>i. measure M2 (?C x i))"
+    finally have "emeasure M2 (Pair x -` Q) = (\<Sum>i. emeasure M2 (?C x i))"
       by simp }
-  ultimately show ?thesis using `Q \<in> sets P` F_sets
-    by (auto intro!: M1.borel_measurable_psuminf M2.Int)
+  ultimately show ?thesis using `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` F_sets
+    by auto
 qed
 
-lemma (in pair_sigma_finite) measure_cut_measurable_snd:
-  assumes "Q \<in> sets P" shows "(\<lambda>y. M1.\<mu> ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
+lemma (in pair_sigma_finite) measurable_emeasure_Pair2:
+  assumes Q: "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"
 proof -
   interpret Q: pair_sigma_finite M2 M1 by default
-  note sets_pair_sigma_algebra_swap[OF assms]
-  from Q.measure_cut_measurable_fst[OF this]
-  show ?thesis by (simp add: vimage_compose[symmetric] comp_def)
-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 (M2 \<Otimes>\<^isub>M 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)
+  have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
+    using Q measurable_pair_swap' by (auto intro: measurable_sets)
+  note Q.measurable_emeasure_Pair1[OF this]
+  moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1)) = (\<lambda>x. (x, y)) -` Q"
+    using Q[THEN sets_into_space] by (auto simp: space_pair_measure)
+  ultimately show ?thesis by simp
 qed
 
-lemma (in pair_sigma_finite) pair_measure_alt:
-  assumes "A \<in> sets P"
-  shows "measure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+ x. measure M2 (Pair x -` A) \<partial>M1)"
-  apply (simp add: pair_measure_def pair_measure_generator_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 :: ereal)"
-    unfolding indicator_def by auto
-  show "(\<integral>\<^isup>+ y. indicator A (x, y) \<partial>M2) = measure M2 (Pair x -` A)"
-    unfolding *
-    apply (subst M2.positive_integral_indicator)
-    apply (rule measurable_cut_fst[OF assms])
-    by simp
+lemma (in pair_sigma_finite) emeasure_pair_measure:
+  assumes "X \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+  shows "emeasure (M1 \<Otimes>\<^isub>M M2) X = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator X (x, y) \<partial>M2 \<partial>M1)" (is "_ = ?\<mu> X")
+proof (rule emeasure_measure_of[OF pair_measure_def])
+  show "positive (sets (M1 \<Otimes>\<^isub>M M2)) ?\<mu>"
+    by (auto simp: positive_def positive_integral_positive)
+  have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y"
+    by (auto simp: indicator_def)
+  show "countably_additive (sets (M1 \<Otimes>\<^isub>M M2)) ?\<mu>"
+  proof (rule countably_additiveI)
+    fix F :: "nat \<Rightarrow> ('a \<times> 'b) set" assume F: "range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2)" "disjoint_family F"
+    from F have *: "\<And>i. F i \<in> sets (M1 \<Otimes>\<^isub>M M2)" "(\<Union>i. F i) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by auto
+    moreover from F have "\<And>i. (\<lambda>x. emeasure M2 (Pair x -` F i)) \<in> borel_measurable M1"
+      by (intro measurable_emeasure_Pair1) auto
+    moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"
+      by (intro disjoint_family_on_bisimulation[OF F(2)]) auto
+    moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M2"
+      using F by (auto simp: sets_Pair1)
+    ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)"
+      by (auto simp add: vimage_UN positive_integral_suminf[symmetric] suminf_emeasure subset_eq emeasure_nonneg sets_Pair1
+               intro!: positive_integral_cong positive_integral_indicator[symmetric])
+  qed
+  show "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2} \<subseteq> Pow (space M1 \<times> space M2)"
+    using space_closed[of M1] space_closed[of M2] by auto
+qed fact
+
+lemma (in pair_sigma_finite) emeasure_pair_measure_alt:
+  assumes X: "X \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+  shows "emeasure (M1  \<Otimes>\<^isub>M M2) X = (\<integral>\<^isup>+x. emeasure M2 (Pair x -` X) \<partial>M1)"
+proof -
+  have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y"
+    by (auto simp: indicator_def)
+  show ?thesis
+    using X by (auto intro!: positive_integral_cong simp: emeasure_pair_measure sets_Pair1)
 qed
 
-lemma (in pair_sigma_finite) pair_measure_times:
-  assumes A: "A \<in> sets M1" and "B \<in> sets M2"
-  shows "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = M1.\<mu> A * measure M2 B"
+lemma (in pair_sigma_finite) emeasure_pair_measure_Times:
+  assumes A: "A \<in> sets M1" and B: "B \<in> sets M2"
+  shows "emeasure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = emeasure M1 A * emeasure M2 B"
 proof -
-  have "measure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = (\<integral>\<^isup>+ x. measure M2 B * indicator A x \<partial>M1)"
-    using assms 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)
+  have "emeasure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = (\<integral>\<^isup>+x. emeasure M2 B * indicator A x \<partial>M1)"
+    using A B by (auto intro!: positive_integral_cong simp: emeasure_pair_measure_alt)
+  also have "\<dots> = emeasure M2 B * emeasure M1 A"
+    using A by (simp add: emeasure_nonneg positive_integral_cmult_indicator)
+  finally show ?thesis
+    by (simp add: ac_simps)
 qed
 
-lemma (in measure_space) measure_not_negative[simp,intro]:
-  assumes A: "A \<in> sets M" shows "\<mu> A \<noteq> - \<infinity>"
-  using positive_measure[OF A] by auto
-
 lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:
-  "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets E \<and> incseq F \<and> (\<Union>i. F i) = space E \<and>
-    (\<forall>i. measure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
+  defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}"
+  shows "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> E \<and> incseq F \<and> (\<Union>i. F i) = space M1 \<times> space M2 \<and>
+    (\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
 proof -
-  obtain F1 :: "nat \<Rightarrow> 'a set" and F2 :: "nat \<Rightarrow> 'b set" where
-    F1: "range F1 \<subseteq> sets M1" "incseq F1" "(\<Union>i. F1 i) = space M1" "\<And>i. M1.\<mu> (F1 i) \<noteq> \<infinity>" and
-    F2: "range F2 \<subseteq> sets M2" "incseq F2" "(\<Union>i. F2 i) = space M2" "\<And>i. M2.\<mu> (F2 i) \<noteq> \<infinity>"
-    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)" by auto
+  from M1.sigma_finite_incseq guess F1 . note F1 = this
+  from M2.sigma_finite_incseq guess F2 . note F2 = this
+  from F1 F2 have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto
   let ?F = "\<lambda>i. F1 i \<times> F2 i"
-  show ?thesis unfolding space_pair_measure
+  show ?thesis
   proof (intro exI[of _ ?F] conjI allI)
-    show "range ?F \<subseteq> sets E" using F1 F2
-      by (fastforce intro!: pair_measure_generatorI)
+    show "range ?F \<subseteq> E" using F1 F2 by (auto simp: E_def) (metis range_subsetD)
   next
     have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"
     proof (intro subsetI)
@@ -448,353 +352,315 @@
         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 E"
-      using space by (auto simp: space pair_measure_generator_def)
+    then show "(\<Union>i. ?F i) = space M1 \<times> space M2"
+      using space by (auto simp: space)
   next
     fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)"
       using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto
   next
     fix i
     from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto
-    with F1 F2 M1.positive_measure[OF this(1)] M2.positive_measure[OF this(2)]
-    show "measure P (F1 i \<times> F2 i) \<noteq> \<infinity>"
-      by (simp add: pair_measure_times)
+    with F1 F2 emeasure_nonneg[of M1 "F1 i"] emeasure_nonneg[of M2 "F2 i"]
+    show "emeasure (M1 \<Otimes>\<^isub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"
+      by (auto simp add: emeasure_pair_measure_Times)
+  qed
+qed
+
+sublocale pair_sigma_finite \<subseteq> sigma_finite_measure "M1 \<Otimes>\<^isub>M M2"
+proof
+  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
+  show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2) \<and> (\<Union>i. F i) = space (M1 \<Otimes>\<^isub>M M2) \<and> (\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)"
+  proof (rule exI[of _ F], intro conjI)
+    show "range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2)" using F by (auto simp: pair_measure_def)
+    show "(\<Union>i. F i) = space (M1 \<Otimes>\<^isub>M M2)"
+      using F by (auto simp: space_pair_measure)
+    show "\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>" using F by auto
   qed
 qed
 
-sublocale pair_sigma_finite \<subseteq> sigma_finite_measure P
-proof
-  show "positive P (measure P)"
-    unfolding pair_measure_def pair_measure_generator_def sigma_def positive_def
-    by (auto intro: M1.positive_integral_positive M2.positive_integral_positive)
-
-  show "countably_additive P (measure P)"
-    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. measure M2 (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
-    ultimately show "(\<Sum>n. measure P (F n)) = measure P (\<Union>i. F i)"
-      by (simp add: pair_measure_alt vimage_UN M1.positive_integral_suminf[symmetric]
-                    M2.measure_countably_additive
-               cong: M1.positive_integral_cong)
-  qed
-
-  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
-  show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets P \<and> (\<Union>i. F i) = space P \<and> (\<forall>i. measure P (F i) \<noteq> \<infinity>)"
-  proof (rule exI[of _ F], intro conjI)
-    show "range F \<subseteq> sets P" using F by (auto simp: pair_measure_def)
-    show "(\<Union>i. F i) = space P"
-      using F by (auto simp: pair_measure_def pair_measure_generator_def)
-    show "\<forall>i. measure P (F i) \<noteq> \<infinity>" using F by auto
-  qed
+lemma sigma_finite_pair_measure:
+  assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B"
+  shows "sigma_finite_measure (A \<Otimes>\<^isub>M B)"
+proof -
+  interpret A: sigma_finite_measure A by fact
+  interpret B: sigma_finite_measure B by fact
+  interpret AB: pair_sigma_finite A  B ..
+  show ?thesis ..
 qed
 
-lemma (in pair_sigma_algebra) sets_swap:
-  assumes "A \<in> sets P"
+lemma sets_pair_swap:
+  assumes "A \<in> sets (M1 \<Otimes>\<^isub>M M2)"
   shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)"
-    (is "_ -` A \<inter> space ?Q \<in> sets ?Q")
-proof -
-  have *: "(\<lambda>(x, y). (y, x)) -` A \<inter> space ?Q = (\<lambda>(x, y). (y, x)) -` A"
-    using `A \<in> sets P` sets_into_space by (auto simp: space_pair_measure)
-  show ?thesis
-    unfolding * using assms by (rule sets_pair_sigma_algebra_swap)
-qed
+  using measurable_pair_swap' assms by (rule measurable_sets)
 
-lemma (in pair_sigma_finite) pair_measure_alt2:
-  assumes A: "A \<in> sets P"
-  shows "\<mu> A = (\<integral>\<^isup>+y. M1.\<mu> ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
-    (is "_ = ?\<nu> A")
+lemma (in pair_sigma_finite) distr_pair_swap:
+  "M1 \<Otimes>\<^isub>M M2 = distr (M2 \<Otimes>\<^isub>M M1) (M1 \<Otimes>\<^isub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D")
 proof -
   interpret Q: pair_sigma_finite M2 M1 by default
   from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
-  have [simp]: "\<And>m. \<lparr> space = space E, sets = sets (sigma E), measure = m \<rparr> = P\<lparr> measure := m \<rparr>"
-    unfolding pair_measure_def by simp
-
-  have "\<mu> A = Q.\<mu> ((\<lambda>(y, x). (x, y)) -` A \<inter> space Q.P)"
-  proof (rule measure_unique_Int_stable_vimage[OF Int_stable_pair_measure_generator])
-    show "measure_space P" "measure_space Q.P" by default
-    show "(\<lambda>(y, x). (x, y)) \<in> measurable Q.P P" by (rule Q.pair_sigma_algebra_swap_measurable)
-    show "sets (sigma E) = sets P" "space E = space P" "A \<in> sets (sigma E)"
-      using assms unfolding pair_measure_def by auto
-    show "range F \<subseteq> sets E" "incseq F" "(\<Union>i. F i) = space E" "\<And>i. \<mu> (F i) \<noteq> \<infinity>"
-      using F `A \<in> sets P` by (auto simp: pair_measure_def)
-    fix X assume "X \<in> sets E"
-    then obtain A B where X[simp]: "X = A \<times> B" and AB: "A \<in> sets M1" "B \<in> sets M2"
-      unfolding pair_measure_def pair_measure_generator_def by auto
-    then have "(\<lambda>(y, x). (x, y)) -` X \<inter> space Q.P = B \<times> A"
-      using M1.sets_into_space M2.sets_into_space by (auto simp: space_pair_measure)
-    then show "\<mu> X = Q.\<mu> ((\<lambda>(y, x). (x, y)) -` X \<inter> space Q.P)"
-      using AB by (simp add: pair_measure_times Q.pair_measure_times ac_simps)
+  let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
+  show ?thesis
+  proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])
+    show "?E \<subseteq> Pow (space ?P)"
+      using space_closed[of M1] space_closed[of M2] by (auto simp: space_pair_measure)
+    show "sets ?P = sigma_sets (space ?P) ?E"
+      by (simp add: sets_pair_measure space_pair_measure)
+    then show "sets ?D = sigma_sets (space ?P) ?E"
+      by simp
+  next
+    show "range F \<subseteq> ?E" "incseq F" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"
+      using F by (auto simp: space_pair_measure)
+  next
+    fix X assume "X \<in> ?E"
+    then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto
+    have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^isub>M M1) = B \<times> A"
+      using sets_into_space[OF A] sets_into_space[OF B] by (auto simp: space_pair_measure)
+    with A B show "emeasure (M1 \<Otimes>\<^isub>M M2) X = emeasure ?D X"
+      by (simp add: emeasure_pair_measure_Times Q.emeasure_pair_measure_Times emeasure_distr
+                    measurable_pair_swap' ac_simps)
   qed
-  then show ?thesis
-    using sets_into_space[OF A] Q.pair_measure_alt[OF sets_swap[OF A]]
-    by (auto simp add: Q.pair_measure_alt space_pair_measure
-             intro!: M2.positive_integral_cong arg_cong[where f="M1.\<mu>"])
 qed
 
-lemma pair_sigma_algebra_sigma:
-  assumes 1: "incseq S1" "(\<Union>i. S1 i) = space E1" "range S1 \<subseteq> sets E1" and E1: "sets E1 \<subseteq> Pow (space E1)"
-  assumes 2: "decseq S2" "(\<Union>i. S2 i) = space E2" "range S2 \<subseteq> sets E2" and E2: "sets E2 \<subseteq> Pow (space E2)"
-  shows "sets (sigma (pair_measure_generator (sigma E1) (sigma E2))) = sets (sigma (pair_measure_generator E1 E2))"
-    (is "sets ?S = sets ?E")
+lemma (in pair_sigma_finite) emeasure_pair_measure_alt2:
+  assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+  shows "emeasure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"
+    (is "_ = ?\<nu> A")
 proof -
-  interpret 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_measure_generator E1 E2) \<subseteq> Pow (space E1 \<times> space E2)"
-    using E1 E2 by (auto simp add: pair_measure_generator_def)
-  interpret E: sigma_algebra ?E unfolding pair_measure_generator_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 pair_measure_generator_def by auto
-    also have "\<dots> = (\<Union>i. A \<times> S2 i)" by auto
-    also have "\<dots> \<in> sets ?E" unfolding pair_measure_generator_def sets_sigma
-      using 2 `A \<in> sets E1`
-      by (intro sigma_sets.Union)
-         (force 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 pair_measure_generator_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_measure_generator_def sets_sigma
-      by (intro sigma_sets.Union)
-         (force 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_measure_generator_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_measure_generator_def)
-    ultimately have "A \<times> B \<in> sets ?E" by auto }
-  then have "sigma_sets (space ?E) (sets (pair_measure_generator (sigma E1) (sigma E2))) \<subseteq> sets ?E"
-    by (intro E.sigma_sets_subset) (auto simp add: pair_measure_generator_def sets_sigma)
-  then have subset: "sets ?S \<subseteq> sets ?E"
-    by (simp add: sets_sigma pair_measure_generator_def)
-  show "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_measure_generator_def sets_sigma intro: sigma_sets.Basic)
-    qed (auto intro: sigma_sets.intros simp: pair_measure_generator_def)
-  next
-    fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
-  qed
+  interpret Q: pair_sigma_finite M2 M1 by default
+  have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1))) = (\<lambda>x. (x, y)) -` A"
+    using sets_into_space[OF A] by (auto simp: space_pair_measure)
+  show ?thesis using A
+    by (subst distr_pair_swap)
+       (simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap']
+                 Q.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A])
 qed
 
 section "Fubinis theorem"
 
 lemma (in pair_sigma_finite) simple_function_cut:
-  assumes f: "simple_function P f" "\<And>x. 0 \<le> f x"
+  assumes f: "simple_function (M1 \<Otimes>\<^isub>M M2) f" "\<And>x. 0 \<le> f x"
   shows "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
+    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
 proof -
-  have f_borel: "f \<in> borel_measurable P"
+  have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
     using f(1) by (rule borel_measurable_simple_function)
-  let ?F = "\<lambda>z. f -` {z} \<inter> space P"
+  let ?F = "\<lambda>z. f -` {z} \<inter> space (M1 \<Otimes>\<^isub>M M2)"
   let ?F' = "\<lambda>x z. Pair x -` ?F z"
   { fix x assume "x \<in> space M1"
     have [simp]: "\<And>z y. indicator (?F z) (x, y) = indicator (?F' x z) y"
       by (auto simp: indicator_def)
-    have "\<And>y. y \<in> space M2 \<Longrightarrow> (x, y) \<in> space P" using `x \<in> space M1`
+    have "\<And>y. y \<in> space M2 \<Longrightarrow> (x, y) \<in> space (M1 \<Otimes>\<^isub>M M2)" using `x \<in> space M1`
       by (simp add: space_pair_measure)
     moreover have "\<And>x z. ?F' x z \<in> sets M2" using f_borel
-      by (intro borel_measurable_vimage measurable_cut_fst)
+      by (rule sets_Pair1[OF measurable_sets]) auto
     ultimately have "simple_function M2 (\<lambda> y. f (x, y))"
-      apply (rule_tac M2.simple_function_cong[THEN iffD2, OF _])
+      apply (rule_tac simple_function_cong[THEN iffD2, OF _])
       apply (rule simple_function_indicator_representation[OF f(1)])
-      using `x \<in> space M1` by (auto simp del: space_sigma) }
+      using `x \<in> space M1` by auto }
   note M2_sf = this
   { fix x assume x: "x \<in> space M1"
-    then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space P. z * M2.\<mu> (?F' x z))"
-      unfolding M2.positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)]
+    then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space (M1 \<Otimes>\<^isub>M M2). z * emeasure M2 (?F' x z))"
+      unfolding positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)]
       unfolding simple_integral_def
     proof (safe intro!: setsum_mono_zero_cong_left)
-      from f(1) show "finite (f ` space P)" by (rule simple_functionD)
+      from f(1) show "finite (f ` space (M1 \<Otimes>\<^isub>M M2))" by (rule simple_functionD)
     next
-      fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space P"
+      fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space (M1 \<Otimes>\<^isub>M M2)"
         using `x \<in> space M1` by (auto simp: space_pair_measure)
     next
-      fix x' y assume "(x', y) \<in> space P"
+      fix x' y assume "(x', y) \<in> space (M1 \<Otimes>\<^isub>M M2)"
         "f (x', y) \<notin> (\<lambda>y. f (x, y)) ` space M2"
       then have *: "?F' x (f (x', y)) = {}"
         by (force simp: space_pair_measure)
-      show  "f (x', y) * M2.\<mu> (?F' x (f (x', y))) = 0"
+      show  "f (x', y) * emeasure M2 (?F' x (f (x', y))) = 0"
         unfolding * by simp
     qed (simp add: vimage_compose[symmetric] comp_def
                    space_pair_measure) }
   note eq = this
-  moreover have "\<And>z. ?F z \<in> sets P"
-    by (auto intro!: f_borel borel_measurable_vimage simp del: space_sigma)
-  moreover then have "\<And>z. (\<lambda>x. M2.\<mu> (?F' x z)) \<in> borel_measurable M1"
-    by (auto intro!: measure_cut_measurable_fst simp del: vimage_Int)
-  moreover have *: "\<And>i x. 0 \<le> M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
-    using f(1)[THEN simple_functionD(2)] f(2) by (intro M2.positive_measure measurable_cut_fst)
-  moreover { fix i assume "i \<in> f`space P"
-    with * have "\<And>x. 0 \<le> i * M2.\<mu> (Pair x -` (f -` {i} \<inter> space P))"
+  moreover have "\<And>z. ?F z \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+    by (auto intro!: f_borel borel_measurable_vimage)
+  moreover then have "\<And>z. (\<lambda>x. emeasure M2 (?F' x z)) \<in> borel_measurable M1"
+    by (auto intro!: measurable_emeasure_Pair1 simp del: vimage_Int)
+  moreover have *: "\<And>i x. 0 \<le> emeasure M2 (Pair x -` (f -` {i} \<inter> space (M1 \<Otimes>\<^isub>M M2)))"
+    using f(1)[THEN simple_functionD(2)] f(2) by (intro emeasure_nonneg)
+  moreover { fix i assume "i \<in> f`space (M1 \<Otimes>\<^isub>M M2)"
+    with * have "\<And>x. 0 \<le> i * emeasure M2 (Pair x -` (f -` {i} \<inter> space (M1 \<Otimes>\<^isub>M M2)))"
       using f(2) by auto }
   ultimately
   show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f" using f(2)
-    by (auto simp del: vimage_Int cong: measurable_cong
-             intro!: M1.borel_measurable_ereal_setsum setsum_cong
-             simp add: M1.positive_integral_setsum simple_integral_def
-                       M1.positive_integral_cmult
-                       M1.positive_integral_cong[OF eq]
+    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" using f(2)
+    by (auto simp del: vimage_Int cong: measurable_cong intro!: setsum_cong
+             simp add: positive_integral_setsum simple_integral_def
+                       positive_integral_cmult
+                       positive_integral_cong[OF eq]
                        positive_integral_eq_simple_integral[OF f]
-                       pair_measure_alt[symmetric])
+                       emeasure_pair_measure_alt[symmetric])
 qed
 
 lemma (in pair_sigma_finite) positive_integral_fst_measurable:
-  assumes f: "f \<in> borel_measurable P"
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
   shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<in> borel_measurable M1"
       (is "?C f \<in> borel_measurable M1")
-    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
+    and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
 proof -
   from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this
-  then have F_borel: "\<And>i. F i \<in> borel_measurable P"
+  then have F_borel: "\<And>i. F i \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
     by (auto intro: borel_measurable_simple_function)
   note sf = simple_function_cut[OF F(1,5)]
   then have "(\<lambda>x. SUP i. ?C (F i) x) \<in> borel_measurable M1"
     using F(1) by auto
   moreover
   { fix x assume "x \<in> space M1"
-    from F measurable_pair_image_snd[OF F_borel`x \<in> space M1`]
+    from F measurable_Pair2[OF F_borel `x \<in> space M1`]
     have "(\<integral>\<^isup>+y. (SUP i. F i (x, y)) \<partial>M2) = (SUP i. ?C (F i) x)"
-      by (intro M2.positive_integral_monotone_convergence_SUP)
+      by (intro positive_integral_monotone_convergence_SUP)
          (auto simp: incseq_Suc_iff le_fun_def)
     then have "(SUP i. ?C (F i) x) = ?C f x"
       unfolding F(4) positive_integral_max_0 by simp }
   note SUPR_C = this
   ultimately show "?C f \<in> borel_measurable M1"
     by (simp cong: measurable_cong)
-  have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>P) = (SUP i. integral\<^isup>P P (F i))"
+  have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>(M1 \<Otimes>\<^isub>M M2)) = (SUP i. integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (F i))"
     using F_borel F
     by (intro positive_integral_monotone_convergence_SUP) auto
-  also have "(SUP i. integral\<^isup>P P (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)"
+  also have "(SUP i. integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)"
     unfolding sf(2) by simp
   also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1" using F sf(1)
-    by (intro M1.positive_integral_monotone_convergence_SUP[symmetric])
-       (auto intro!: M2.positive_integral_mono M2.positive_integral_positive
-                simp: incseq_Suc_iff le_fun_def)
+    by (intro positive_integral_monotone_convergence_SUP[symmetric])
+       (auto intro!: positive_integral_mono positive_integral_positive
+             simp: incseq_Suc_iff le_fun_def)
   also have "\<dots> = \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. (SUP i. F i (x, y)) \<partial>M2) \<partial>M1"
     using F_borel F(2,5)
-    by (auto intro!: M1.positive_integral_cong M2.positive_integral_monotone_convergence_SUP[symmetric]
-             simp: incseq_Suc_iff le_fun_def measurable_pair_image_snd)
-  finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P P f"
+    by (auto intro!: positive_integral_cong positive_integral_monotone_convergence_SUP[symmetric] measurable_Pair2
+             simp: incseq_Suc_iff le_fun_def)
+  finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
     using F by (simp add: positive_integral_max_0)
 qed
 
-lemma (in pair_sigma_finite) measure_preserving_swap:
-  "(\<lambda>(x,y). (y, x)) \<in> measure_preserving (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
-proof
-  interpret Q: pair_sigma_finite M2 M1 by default
-  show *: "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)"
-    using pair_sigma_algebra_swap_measurable .
-  fix X assume "X \<in> sets (M2 \<Otimes>\<^isub>M M1)"
-  from measurable_sets[OF * this] this Q.sets_into_space[OF this]
-  show "measure (M1 \<Otimes>\<^isub>M M2) ((\<lambda>(x, y). (y, x)) -` X \<inter> space P) = measure (M2 \<Otimes>\<^isub>M M1) X"
-    by (auto intro!: M1.positive_integral_cong arg_cong[where f="M2.\<mu>"]
-      simp: pair_measure_alt Q.pair_measure_alt2 space_pair_measure)
-qed
-
-lemma (in pair_sigma_finite) positive_integral_product_swap:
-  assumes f: "f \<in> borel_measurable P"
-  shows "(\<integral>\<^isup>+x. f (case x of (x,y)\<Rightarrow>(y,x)) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P P f"
+lemma (in pair_sigma_finite) positive_integral_snd_measurable:
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
 proof -
   interpret Q: pair_sigma_finite M2 M1 by default
-  have "sigma_algebra P" by default
-  with f show ?thesis
-    by (subst Q.positive_integral_vimage[OF _ Q.measure_preserving_swap]) auto
-qed
-
-lemma (in pair_sigma_finite) positive_integral_snd_measurable:
-  assumes f: "f \<in> borel_measurable P"
-  shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P P f"
-proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
-  note pair_sigma_algebra_measurable[OF f]
+  note measurable_pair_swap[OF f]
   from Q.positive_integral_fst_measurable[OF this]
-  have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P)"
+  have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1))"
     by simp
-  also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>Q.P) = integral\<^isup>P P f"
-    unfolding positive_integral_product_swap[OF f, symmetric]
-    by (auto intro!: Q.positive_integral_cong)
+  also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f"
+    by (subst distr_pair_swap)
+       (auto simp: positive_integral_distr[OF measurable_pair_swap' f] intro!: positive_integral_cong)
   finally show ?thesis .
 qed
 
 lemma (in pair_sigma_finite) Fubini:
-  assumes f: "f \<in> borel_measurable P"
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
   shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
   unfolding positive_integral_snd_measurable[OF assms]
   unfolding positive_integral_fst_measurable[OF assms] ..
 
 lemma (in pair_sigma_finite) AE_pair:
-  assumes "AE x in P. Q x"
+  assumes "AE x in (M1 \<Otimes>\<^isub>M M2). Q x"
   shows "AE x in M1. (AE y in M2. Q (x, y))"
 proof -
-  obtain N where N: "N \<in> sets P" "\<mu> N = 0" "{x\<in>space P. \<not> Q x} \<subseteq> N"
-    using assms unfolding almost_everywhere_def by auto
+  obtain N where N: "N \<in> sets (M1 \<Otimes>\<^isub>M M2)" "emeasure (M1 \<Otimes>\<^isub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^isub>M M2). \<not> Q x} \<subseteq> N"
+    using assms unfolding eventually_ae_filter by auto
   show ?thesis
-  proof (rule M1.AE_I)
-    from N measure_cut_measurable_fst[OF `N \<in> sets P`]
-    show "M1.\<mu> {x\<in>space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} = 0"
-      by (auto simp: pair_measure_alt M1.positive_integral_0_iff)
-    show "{x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0} \<in> sets M1"
-      by (intro M1.borel_measurable_ereal_neq_const measure_cut_measurable_fst N)
-    { fix x assume "x \<in> space M1" "M2.\<mu> (Pair x -` N) = 0"
-      have "M2.almost_everywhere (\<lambda>y. Q (x, y))"
-      proof (rule M2.AE_I)
-        show "M2.\<mu> (Pair x -` N) = 0" by fact
-        show "Pair x -` N \<in> sets M2" by (intro measurable_cut_fst N)
+  proof (rule AE_I)
+    from N measurable_emeasure_Pair1[OF `N \<in> sets (M1 \<Otimes>\<^isub>M M2)`]
+    show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
+      by (auto simp: emeasure_pair_measure_alt positive_integral_0_iff emeasure_nonneg)
+    show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
+      by (intro borel_measurable_ereal_neq_const measurable_emeasure_Pair1 N)
+    { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0"
+      have "AE y in M2. Q (x, y)"
+      proof (rule AE_I)
+        show "emeasure M2 (Pair x -` N) = 0" by fact
+        show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1)
         show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
-          using N `x \<in> space M1` unfolding space_sigma space_pair_measure by auto
+          using N `x \<in> space M1` unfolding space_pair_measure by auto
       qed }
-    then show "{x \<in> space M1. \<not> M2.almost_everywhere (\<lambda>y. Q (x, y))} \<subseteq> {x \<in> space M1. M2.\<mu> (Pair x -` N) \<noteq> 0}"
+    then show "{x \<in> space M1. \<not> (AE y in M2. Q (x, y))} \<subseteq> {x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0}"
       by auto
   qed
 qed
 
-lemma (in pair_sigma_algebra) measurable_product_swap:
-  "f \<in> measurable (M2 \<Otimes>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable P M"
+lemma (in pair_sigma_finite) AE_pair_measure:
+  assumes "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P x} \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+  assumes ae: "AE x in M1. AE y in M2. P (x, y)"
+  shows "AE x in M1 \<Otimes>\<^isub>M M2. P x"
+proof (subst AE_iff_measurable[OF _ refl])
+  show "{x\<in>space (M1 \<Otimes>\<^isub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+    by (rule sets_Collect) fact
+  then have "emeasure (M1 \<Otimes>\<^isub>M M2) {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} =
+      (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)"
+    by (simp add: emeasure_pair_measure)
+  also have "\<dots> = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. 0 \<partial>M2 \<partial>M1)"
+    using ae
+    apply (safe intro!: positive_integral_cong_AE)
+    apply (intro AE_I2)
+    apply (safe intro!: positive_integral_cong_AE)
+    apply auto
+    done
+  finally show "emeasure (M1 \<Otimes>\<^isub>M M2) {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} = 0" by simp
+qed
+
+lemma (in pair_sigma_finite) AE_pair_iff:
+  "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow>
+    (AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x))"
+  using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto
+
+lemma AE_distr_iff:
+  assumes f: "f \<in> measurable M N" and P: "{x \<in> space N. P x} \<in> sets N"
+  shows "(AE x in distr M N f. P x) \<longleftrightarrow> (AE x in M. P (f x))"
+proof (subst (1 2) AE_iff_measurable[OF _ refl])
+  from P show "{x \<in> space (distr M N f). \<not> P x} \<in> sets (distr M N f)"
+    by (auto intro!: sets_Collect_neg)
+  moreover
+  have "f -` {x \<in> space N. P x} \<inter> space M = {x \<in> space M. P (f x)}"
+    using f by (auto dest: measurable_space)
+  then show "{x \<in> space M. \<not> P (f x)} \<in> sets M"
+    using measurable_sets[OF f P] by (auto intro!: sets_Collect_neg)
+  moreover have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}"
+    using f by (auto dest: measurable_space)
+  ultimately show "(emeasure (distr M N f) {x \<in> space (distr M N f). \<not> P x} = 0) =
+    (emeasure M {x \<in> space M. \<not> P (f x)} = 0)"
+    using f by (simp add: emeasure_distr)
+qed
+
+lemma (in pair_sigma_finite) AE_commute:
+  assumes P: "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^isub>M M2)"
+  shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)"
 proof -
-  interpret Q: pair_sigma_algebra M2 M1 by default
-  show ?thesis
-    using pair_sigma_algebra_measurable[of "\<lambda>(x,y). f (y, x)"]
-    by (auto intro!: pair_sigma_algebra_measurable Q.pair_sigma_algebra_measurable iffI)
+  interpret Q: pair_sigma_finite M2 M1 ..
+  have [simp]: "\<And>x. (fst (case x of (x, y) \<Rightarrow> (y, x))) = snd x" "\<And>x. (snd (case x of (x, y) \<Rightarrow> (y, x))) = fst x"
+    by auto
+  have "{x \<in> space (M2 \<Otimes>\<^isub>M M1). P (snd x) (fst x)} =
+    (\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^isub>M M1)"
+    by (auto simp: space_pair_measure)
+  also have "\<dots> \<in> sets (M2 \<Otimes>\<^isub>M M1)"
+    by (intro sets_pair_swap P)
+  finally show ?thesis
+    apply (subst AE_pair_iff[OF P])
+    apply (subst distr_pair_swap)
+    apply (subst AE_distr_iff[OF measurable_pair_swap' P])
+    apply (subst Q.AE_pair_iff)
+    apply simp_all
+    done
 qed
 
 lemma (in pair_sigma_finite) integrable_product_swap:
-  assumes "integrable P f"
+  assumes "integrable (M1 \<Otimes>\<^isub>M M2) f"
   shows "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x))"
 proof -
   interpret Q: pair_sigma_finite M2 M1 by default
   have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
   show ?thesis unfolding *
-    using assms unfolding integrable_def
-    apply (subst (1 2) positive_integral_product_swap)
-    using `integrable P f` unfolding integrable_def
-    by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
+    by (rule integrable_distr[OF measurable_pair_swap'])
+       (simp add: distr_pair_swap[symmetric] assms)
 qed
 
 lemma (in pair_sigma_finite) integrable_product_swap_iff:
-  "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable P f"
+  "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^isub>M M2) f"
 proof -
   interpret Q: pair_sigma_finite M2 M1 by default
   from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f]
@@ -802,27 +668,25 @@
 qed
 
 lemma (in pair_sigma_finite) integral_product_swap:
-  assumes "integrable P f"
-  shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>L P f"
+  assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+  shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f"
 proof -
-  interpret Q: pair_sigma_finite M2 M1 by default
   have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)
-  show ?thesis
-    unfolding lebesgue_integral_def *
-    apply (subst (1 2) positive_integral_product_swap)
-    using `integrable P f` unfolding integrable_def
-    by (auto simp: *[symmetric] Q.measurable_product_swap[symmetric])
+  show ?thesis unfolding *
+    by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric])
 qed
 
 lemma (in pair_sigma_finite) integrable_fst_measurable:
-  assumes f: "integrable P f"
-  shows "M1.almost_everywhere (\<lambda>x. integrable M2 (\<lambda> y. f (x, y)))" (is "?AE")
-    and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L P f" (is "?INT")
+  assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f"
+  shows "AE x in M1. integrable M2 (\<lambda> y. f (x, y))" (is "?AE")
+    and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f" (is "?INT")
 proof -
+  have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+    using f by auto
   let ?pf = "\<lambda>x. ereal (f x)" and ?nf = "\<lambda>x. ereal (- f x)"
   have
-    borel: "?nf \<in> borel_measurable P""?pf \<in> borel_measurable P" and
-    int: "integral\<^isup>P P ?nf \<noteq> \<infinity>" "integral\<^isup>P P ?pf \<noteq> \<infinity>"
+    borel: "?nf \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)""?pf \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" and
+    int: "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) ?nf \<noteq> \<infinity>" "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) ?pf \<noteq> \<infinity>"
     using assms by auto
   have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
      "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"
@@ -831,69 +695,92 @@
   with borel[THEN positive_integral_fst_measurable(1)]
   have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
     "AE x in M1. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"
-    by (auto intro!: M1.positive_integral_PInf_AE )
+    by (auto intro!: positive_integral_PInf_AE )
   then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
     "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"
-    by (auto simp: M2.positive_integral_positive)
+    by (auto simp: positive_integral_positive)
   from AE_pos show ?AE using assms
-    by (simp add: measurable_pair_image_snd integrable_def)
+    by (simp add: measurable_Pair2[OF f_borel] integrable_def)
   { fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
-      using M2.positive_integral_positive
-      by (intro M1.positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder)
+      using positive_integral_positive
+      by (intro positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder)
     then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }
   note this[simp]
-  { fix f assume borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable P"
-      and int: "integral\<^isup>P P (\<lambda>x. ereal (f x)) \<noteq> \<infinity>"
-      and AE: "M1.almost_everywhere (\<lambda>x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>)"
+  { fix f assume borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+      and int: "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (\<lambda>x. ereal (f x)) \<noteq> \<infinity>"
+      and AE: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"
     have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")
     proof (intro integrable_def[THEN iffD2] conjI)
       show "?f \<in> borel_measurable M1"
-        using borel by (auto intro!: M1.borel_measurable_real_of_ereal positive_integral_fst_measurable)
+        using borel by (auto intro!: positive_integral_fst_measurable)
       have "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y))  \<partial>M2) \<partial>M1)"
-        using AE M2.positive_integral_positive
-        by (auto intro!: M1.positive_integral_cong_AE simp: ereal_real)
+        using AE positive_integral_positive[of M2]
+        by (auto intro!: positive_integral_cong_AE simp: ereal_real)
       then show "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) \<noteq> \<infinity>"
         using positive_integral_fst_measurable[OF borel] int by simp
       have "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)"
-        by (intro M1.positive_integral_cong_pos)
-           (simp add: M2.positive_integral_positive real_of_ereal_pos)
+        by (intro positive_integral_cong_pos)
+           (simp add: positive_integral_positive real_of_ereal_pos)
       then show "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp
     qed }
   with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]
   show ?INT
-    unfolding lebesgue_integral_def[of P] lebesgue_integral_def[of M2]
+    unfolding lebesgue_integral_def[of "M1 \<Otimes>\<^isub>M M2"] lebesgue_integral_def[of M2]
       borel[THEN positive_integral_fst_measurable(2), symmetric]
-    using AE[THEN M1.integral_real]
+    using AE[THEN integral_real]
     by simp
 qed
 
 lemma (in pair_sigma_finite) integrable_snd_measurable:
-  assumes f: "integrable P f"
-  shows "M2.almost_everywhere (\<lambda>y. integrable M1 (\<lambda>x. f (x, y)))" (is "?AE")
-    and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>L P f" (is "?INT")
+  assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f"
+  shows "AE y in M2. integrable M1 (\<lambda>x. f (x, y))" (is "?AE")
+    and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f" (is "?INT")
 proof -
   interpret Q: pair_sigma_finite M2 M1 by default
-  have Q_int: "integrable Q.P (\<lambda>(x, y). f (y, x))"
+  have Q_int: "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x, y). f (y, x))"
     using f unfolding integrable_product_swap_iff .
   show ?INT
     using Q.integrable_fst_measurable(2)[OF Q_int]
-    using integral_product_swap[OF f] by simp
+    using integral_product_swap[of f] f by auto
   show ?AE
     using Q.integrable_fst_measurable(1)[OF Q_int]
     by simp
 qed
 
+lemma (in pair_sigma_finite) positive_integral_fst_measurable':
+  assumes f: "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+  shows "(\<lambda>x. \<integral>\<^isup>+ y. f x y \<partial>M2) \<in> borel_measurable M1"
+  using positive_integral_fst_measurable(1)[OF f] by simp
+
+lemma (in pair_sigma_finite) integral_fst_measurable:
+  "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow> (\<lambda>x. \<integral> y. f x y \<partial>M2) \<in> borel_measurable M1"
+  by (auto simp: lebesgue_integral_def intro!: borel_measurable_diff positive_integral_fst_measurable')
+
+lemma (in pair_sigma_finite) positive_integral_snd_measurable':
+  assumes f: "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+  shows "(\<lambda>y. \<integral>\<^isup>+ x. f x y \<partial>M1) \<in> borel_measurable M2"
+proof -
+  interpret Q: pair_sigma_finite M2 M1 ..
+  show ?thesis
+    using measurable_pair_swap[OF f]
+    by (intro Q.positive_integral_fst_measurable') (simp add: split_beta')
+qed
+
+lemma (in pair_sigma_finite) integral_snd_measurable:
+  "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow> (\<lambda>y. \<integral> x. f x y \<partial>M1) \<in> borel_measurable M2"
+  by (auto simp: lebesgue_integral_def intro!: borel_measurable_diff positive_integral_snd_measurable')
+
 lemma (in pair_sigma_finite) Fubini_integral:
-  assumes f: "integrable P f"
+  assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f"
   shows "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1)"
   unfolding integrable_snd_measurable[OF assms]
   unfolding integrable_fst_measurable[OF assms] ..
 
-section "Products on finite spaces"
+section {* Products on counting spaces, densities and distributions *}
 
 lemma sigma_sets_pair_measure_generator_finite:
   assumes "finite A" and "finite B"
-  shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<in> Pow A \<and> b \<in> Pow B} = Pow (A \<times> B)"
+  shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<subseteq> A \<and> b \<subseteq> B} = Pow (A \<times> B)"
   (is "sigma_sets ?prod ?sets = _")
 proof safe
   have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)
@@ -904,8 +791,7 @@
     case empty show ?case by (rule sigma_sets.Empty)
   next
     case (insert a x)
-    hence "{a} \<in> sigma_sets ?prod ?sets"
-      by (auto simp: pair_measure_generator_def intro!: sigma_sets.Basic)
+    hence "{a} \<in> sigma_sets ?prod ?sets" by auto
     moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto
     ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)
   qed
@@ -916,48 +802,142 @@
   show "a \<in> A" and "b \<in> B" by auto
 qed
 
-locale pair_finite_sigma_algebra = pair_sigma_algebra M1 M2 + M1: finite_sigma_algebra M1 + M2: finite_sigma_algebra M2 for M1 M2
-
-lemma (in pair_finite_sigma_algebra) finite_pair_sigma_algebra:
-  shows "P = \<lparr> space = space M1 \<times> space M2, sets = Pow (space M1 \<times> space M2), \<dots> = algebra.more P \<rparr>"
-proof -
-  show ?thesis
-    using sigma_sets_pair_measure_generator_finite[OF M1.finite_space M2.finite_space]
-    by (intro algebra.equality) (simp_all add: pair_measure_def pair_measure_generator_def sigma_def)
-qed
-
-sublocale pair_finite_sigma_algebra \<subseteq> finite_sigma_algebra P
-proof
-  show "finite (space P)"
-    using M1.finite_space M2.finite_space
-    by (subst finite_pair_sigma_algebra) simp
-  show "sets P = Pow (space P)"
-    by (subst (1 2) finite_pair_sigma_algebra) simp
+lemma pair_measure_count_space:
+  assumes A: "finite A" and B: "finite B"
+  shows "count_space A \<Otimes>\<^isub>M count_space B = count_space (A \<times> B)" (is "?P = ?C")
+proof (rule measure_eqI)
+  interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact
+  interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact
+  interpret P: pair_sigma_finite "count_space A" "count_space B" by default
+  show eq: "sets ?P = sets ?C"
+    by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B)
+  fix X assume X: "X \<in> sets ?P"
+  with eq have X_subset: "X \<subseteq> A \<times> B" by simp
+  with A B have fin_Pair: "\<And>x. finite (Pair x -` X)"
+    by (intro finite_subset[OF _ B]) auto
+  have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B)
+  show "emeasure ?P X = emeasure ?C X"
+    apply (subst P.emeasure_pair_measure_alt[OF X])
+    apply (subst emeasure_count_space)
+    using X_subset apply auto []
+    apply (simp add: fin_Pair emeasure_count_space X_subset fin_X)
+    apply (subst positive_integral_count_space)
+    using A apply simp
+    apply (simp del: real_of_nat_setsum add: real_of_nat_setsum[symmetric])
+    apply (subst card_gt_0_iff)
+    apply (simp add: fin_Pair)
+    apply (subst card_SigmaI[symmetric])
+    using A apply simp
+    using fin_Pair apply simp
+    using X_subset apply (auto intro!: arg_cong[where f=card])
+    done
 qed
 
-locale pair_finite_space = pair_sigma_finite M1 M2 + pair_finite_sigma_algebra M1 M2 +
-  M1: finite_measure_space M1 + M2: finite_measure_space M2 for M1 M2
+lemma pair_measure_density:
+  assumes f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x"
+  assumes g: "g \<in> borel_measurable M2" "AE x in M2. 0 \<le> g x"
+  assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
+  assumes "sigma_finite_measure (density M1 f)" "sigma_finite_measure (density M2 g)"
+  shows "density M1 f \<Otimes>\<^isub>M density M2 g = density (M1 \<Otimes>\<^isub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R")
+proof (rule measure_eqI)
+  interpret M1: sigma_finite_measure M1 by fact
+  interpret M2: sigma_finite_measure M2 by fact
+  interpret D1: sigma_finite_measure "density M1 f" by fact
+  interpret D2: sigma_finite_measure "density M2 g" by fact
+  interpret L: pair_sigma_finite "density M1 f" "density M2 g" ..
+  interpret R: pair_sigma_finite M1 M2 ..
+
+  fix A assume A: "A \<in> sets ?L"
+  then have indicator_eq: "\<And>x y. indicator A (x, y) = indicator (Pair x -` A) y"
+   and Pair_A: "\<And>x. Pair x -` A \<in> sets M2"
+    by (auto simp: indicator_def sets_Pair1)
+  have f_fst: "(\<lambda>p. f (fst p)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+    using measurable_comp[OF measurable_fst f(1)] by (simp add: comp_def)
+  have g_snd: "(\<lambda>p. g (snd p)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)"
+    using measurable_comp[OF measurable_snd g(1)] by (simp add: comp_def)
+  have "(\<lambda>x. \<integral>\<^isup>+ y. g (snd (x, y)) * indicator A (x, y) \<partial>M2) \<in> borel_measurable M1"
+    using g_snd Pair_A A by (intro R.positive_integral_fst_measurable) auto
+  then have int_g: "(\<lambda>x. \<integral>\<^isup>+ y. g y * indicator A (x, y) \<partial>M2) \<in> borel_measurable M1"
+    by simp
 
-lemma (in pair_finite_space) pair_measure_Pair[simp]:
-  assumes "a \<in> space M1" "b \<in> space M2"
-  shows "\<mu> {(a, b)} = M1.\<mu> {a} * M2.\<mu> {b}"
+  show "emeasure ?L A = emeasure ?R A"
+    apply (subst L.emeasure_pair_measure[OF A])
+    apply (subst emeasure_density)
+        using f_fst g_snd apply (simp add: split_beta')
+      using A apply simp
+    apply (subst positive_integral_density[OF g])
+      apply (simp add: indicator_eq Pair_A)
+    apply (subst positive_integral_density[OF f])
+      apply (rule int_g)
+    apply (subst R.positive_integral_fst_measurable(2)[symmetric])
+      using f g A Pair_A f_fst g_snd
+      apply (auto intro!: positive_integral_cong_AE R.measurable_emeasure_Pair1
+                  simp: positive_integral_cmult indicator_eq split_beta')
+    apply (intro AE_I2 impI)
+    apply (subst mult_assoc)
+    apply (subst positive_integral_cmult)
+          apply auto
+    done
+qed simp
+
+lemma sigma_finite_measure_distr:
+  assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N"
+  shows "sigma_finite_measure M"
 proof -
-  have "\<mu> ({a}\<times>{b}) = M1.\<mu> {a} * M2.\<mu> {b}"
-    using M1.sets_eq_Pow M2.sets_eq_Pow assms
-    by (subst pair_measure_times) auto
-  then show ?thesis by simp
+  interpret sigma_finite_measure "distr M N f" by fact
+  from sigma_finite_disjoint guess A . note A = this
+  show ?thesis
+  proof (unfold_locales, intro conjI exI allI)
+    show "range (\<lambda>i. f -` A i \<inter> space M) \<subseteq> sets M"
+      using A f by (auto intro!: measurable_sets)
+    show "(\<Union>i. f -` A i \<inter> space M) = space M"
+      using A(1) A(2)[symmetric] f by (auto simp: measurable_def Pi_def)
+    fix i show "emeasure M (f -` A i \<inter> space M) \<noteq> \<infinity>"
+      using f A(1,2) A(3)[of i] by (simp add: emeasure_distr subset_eq)
+  qed
 qed
 
-lemma (in pair_finite_space) pair_measure_singleton[simp]:
-  assumes "x \<in> space M1 \<times> space M2"
-  shows "\<mu> {x} = M1.\<mu> {fst x} * M2.\<mu> {snd x}"
-  using pair_measure_Pair assms by (cases x) auto
+lemma measurable_cong':
+  assumes sets: "sets M = sets M'" "sets N = sets N'"
+  shows "measurable M N = measurable M' N'"
+  using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def)
 
-sublocale pair_finite_space \<subseteq> finite_measure_space P
-proof unfold_locales
-  show "measure P (space P) \<noteq> \<infinity>"
-    by (subst (2) finite_pair_sigma_algebra)
-       (simp add: pair_measure_times)
+lemma pair_measure_distr:
+  assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T"
+  assumes "sigma_finite_measure (distr M S f)" "sigma_finite_measure (distr N T g)"
+  shows "distr M S f \<Otimes>\<^isub>M distr N T g = distr (M \<Otimes>\<^isub>M N) (S \<Otimes>\<^isub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D")
+proof (rule measure_eqI)
+  show "sets ?P = sets ?D"
+    by simp
+  interpret S: sigma_finite_measure "distr M S f" by fact
+  interpret T: sigma_finite_measure "distr N T g" by fact
+  interpret ST: pair_sigma_finite "distr M S f"  "distr N T g" ..
+  interpret M: sigma_finite_measure M by (rule sigma_finite_measure_distr) fact+
+  interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+
+  interpret MN: pair_sigma_finite M N ..
+  interpret SN: pair_sigma_finite "distr M S f" N ..
+  have [simp]: 
+    "\<And>f g. fst \<circ> (\<lambda>(x, y). (f x, g y)) = f \<circ> fst" "\<And>f g. snd \<circ> (\<lambda>(x, y). (f x, g y)) = g \<circ> snd"
+    by auto
+  then have fg: "(\<lambda>(x, y). (f x, g y)) \<in> measurable (M \<Otimes>\<^isub>M N) (S \<Otimes>\<^isub>M T)"
+    using measurable_comp[OF measurable_fst f] measurable_comp[OF measurable_snd g]
+    by (auto simp: measurable_pair_iff)
+  fix A assume A: "A \<in> sets ?P"
+  then have "emeasure ?P A = (\<integral>\<^isup>+x. emeasure (distr N T g) (Pair x -` A) \<partial>distr M S f)"
+    by (rule ST.emeasure_pair_measure_alt)
+  also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (g -` (Pair x -` A) \<inter> space N) \<partial>distr M S f)"
+    using g A by (simp add: sets_Pair1 emeasure_distr)
+  also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (g -` (Pair (f x) -` A) \<inter> space N) \<partial>M)"
+    using f g A ST.measurable_emeasure_Pair1[OF A]
+    by (intro positive_integral_distr) (auto simp add: sets_Pair1 emeasure_distr)
+  also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (Pair x -` ((\<lambda>(x, y). (f x, g y)) -` A \<inter> space (M \<Otimes>\<^isub>M N))) \<partial>M)"
+    by (intro positive_integral_cong arg_cong2[where f=emeasure]) (auto simp: space_pair_measure)
+  also have "\<dots> = emeasure (M \<Otimes>\<^isub>M N) ((\<lambda>(x, y). (f x, g y)) -` A \<inter> space (M \<Otimes>\<^isub>M N))"
+    using fg by (intro MN.emeasure_pair_measure_alt[symmetric] measurable_sets[OF _ A])
+                (auto cong: measurable_cong')
+  also have "\<dots> = emeasure ?D A"
+    using fg A by (subst emeasure_distr) auto
+  finally show "emeasure ?P A = emeasure ?D A" .
 qed
 
 end
\ No newline at end of file
--- a/src/HOL/Probability/Borel_Space.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Borel_Space.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -11,16 +11,14 @@
 
 section "Generic Borel spaces"
 
-definition "borel = sigma \<lparr> space = UNIV::'a::topological_space set, sets = {S. open S}\<rparr>"
-abbreviation "borel_measurable M \<equiv> measurable M borel"
+definition borel :: "'a::topological_space measure" where
+  "borel = sigma UNIV {S. open S}"
 
-interpretation borel: sigma_algebra borel
-  by (auto simp: borel_def intro!: sigma_algebra_sigma)
+abbreviation "borel_measurable M \<equiv> measurable M borel"
 
 lemma in_borel_measurable:
    "f \<in> borel_measurable M \<longleftrightarrow>
-    (\<forall>S \<in> sets (sigma \<lparr> space = UNIV, sets = {S. open S}\<rparr>).
-      f -` S \<inter> space M \<in> sets M)"
+    (\<forall>S \<in> sigma_sets UNIV {S. open S}. f -` S \<inter> space M \<in> sets M)"
   by (auto simp add: measurable_def borel_def)
 
 lemma in_borel_measurable_borel:
@@ -36,7 +34,7 @@
   assumes "open A" shows "A \<in> sets borel"
 proof -
   have "A \<in> {S. open S}" unfolding mem_Collect_eq using assms .
-  thus ?thesis unfolding borel_def sigma_def by (auto intro!: sigma_sets.Basic)
+  thus ?thesis unfolding borel_def by auto
 qed
 
 lemma borel_closed[simp]:
@@ -48,9 +46,9 @@
 qed
 
 lemma borel_comp[intro,simp]: "A \<in> sets borel \<Longrightarrow> - A \<in> sets borel"
-  unfolding Compl_eq_Diff_UNIV by (intro borel.Diff) auto
+  unfolding Compl_eq_Diff_UNIV by (intro Diff) auto
 
-lemma (in sigma_algebra) borel_measurable_vimage:
+lemma 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"
@@ -65,12 +63,12 @@
   thus ?thesis by auto
 qed
 
-lemma (in sigma_algebra) borel_measurableI:
+lemma 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)
+proof (rule measurable_measure_of, simp_all)
   fix S :: "'x set" assume "open S" thus "f -` S \<inter> space M \<in> sets M"
     using assms[of S] by simp
 qed
@@ -78,22 +76,22 @@
 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)
+  proof (rule insert_in_sets)
     show "{x} \<in> sets borel"
       using closed_singleton[of x] by (rule borel_closed)
   qed simp
 
-lemma (in sigma_algebra) borel_measurable_const[simp, intro]:
+lemma borel_measurable_const[simp, intro]:
   "(\<lambda>x. c) \<in> borel_measurable M"
-  by (auto intro!: measurable_const)
+  by auto
 
-lemma (in sigma_algebra) borel_measurable_indicator[simp, intro!]:
+lemma borel_measurable_indicator[simp, intro!]:
   assumes A: "A \<in> sets M"
   shows "indicator A \<in> borel_measurable M"
   unfolding indicator_def [abs_def] using A
-  by (auto intro!: measurable_If_set borel_measurable_const)
+  by (auto intro!: measurable_If_set)
 
-lemma (in sigma_algebra) borel_measurable_indicator_iff:
+lemma borel_measurable_indicator_iff:
   "(indicator A :: 'a \<Rightarrow> 'x::{t1_space, zero_neq_one}) \<in> borel_measurable M \<longleftrightarrow> A \<inter> space M \<in> sets M"
     (is "?I \<in> borel_measurable M \<longleftrightarrow> _")
 proof
@@ -111,49 +109,7 @@
   ultimately show "?I \<in> borel_measurable M" by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_restricted:
-  fixes f :: "'a \<Rightarrow> ereal" 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 :: "ereal 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 fastforce
-    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 :: "ereal 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:
+lemma borel_measurable_subalgebra:
   assumes "sets N \<subseteq> sets M" "space N = space M" "f \<in> borel_measurable N"
   shows "f \<in> borel_measurable M"
   using assms unfolding measurable_def by auto
@@ -220,7 +176,7 @@
   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]:
+lemma borel_measurable_less[simp, intro]:
   fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -233,7 +189,7 @@
     by simp (blast intro: measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_le[simp, intro]:
+lemma borel_measurable_le[simp, intro]:
   fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -245,7 +201,7 @@
     by simp blast
 qed
 
-lemma (in sigma_algebra) borel_measurable_eq[simp, intro]:
+lemma borel_measurable_eq[simp, intro]:
   fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -257,7 +213,7 @@
   thus ?thesis using f g by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_neq[simp, intro]:
+lemma borel_measurable_neq[simp, intro]:
   fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -351,23 +307,70 @@
   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 borel_sigma_sets_subset:
+  "A \<subseteq> sets borel \<Longrightarrow> sigma_sets UNIV A \<subseteq> sets borel"
+  using sigma_sets_subset[of A borel] by simp
+
+lemma borel_eq_sigmaI1:
+  fixes F :: "'i \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
+  assumes borel_eq: "borel = sigma UNIV X"
+  assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (range F))"
+  assumes F: "\<And>i. F i \<in> sets borel"
+  shows "borel = sigma UNIV (range F)"
+  unfolding borel_def
+proof (intro sigma_eqI antisym)
+  have borel_rev_eq: "sigma_sets UNIV {S::'a set. open S} = sets borel"
+    unfolding borel_def by simp
+  also have "\<dots> = sigma_sets UNIV X"
+    unfolding borel_eq by simp
+  also have "\<dots> \<subseteq> sigma_sets UNIV (range F)"
+    using X by (intro sigma_algebra.sigma_sets_subset[OF sigma_algebra_sigma_sets]) auto
+  finally show "sigma_sets UNIV {S. open S} \<subseteq> sigma_sets UNIV (range F)" .
+  show "sigma_sets UNIV (range F) \<subseteq> sigma_sets UNIV {S. open S}"
+    unfolding borel_rev_eq using F by (intro borel_sigma_sets_subset) auto
+qed auto
 
-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 borel_eq_sigmaI2:
+  fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set"
+    and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
+  assumes borel_eq: "borel = sigma UNIV (range (\<lambda>(i, j). G i j))"
+  assumes X: "\<And>i j. G i j \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
+  assumes F: "\<And>i j. F i j \<in> sets borel"
+  shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
+  using assms by (intro borel_eq_sigmaI1[where X="range (\<lambda>(i, j). G i j)" and F="(\<lambda>(i, j). F i j)"]) auto
+
+lemma borel_eq_sigmaI3:
+  fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and X :: "'a::topological_space set set"
+  assumes borel_eq: "borel = sigma UNIV X"
+  assumes X: "\<And>x. x \<in> X \<Longrightarrow> x \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
+  assumes F: "\<And>i j. F i j \<in> sets borel"
+  shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
+  using assms by (intro borel_eq_sigmaI1[where X=X and F="(\<lambda>(i, j). F i j)"]) auto
+
+lemma borel_eq_sigmaI4:
+  fixes F :: "'i \<Rightarrow> 'a::topological_space set"
+    and G :: "'l \<Rightarrow> 'k \<Rightarrow> 'a::topological_space set"
+  assumes borel_eq: "borel = sigma UNIV (range (\<lambda>(i, j). G i j))"
+  assumes X: "\<And>i j. G i j \<in> sets (sigma UNIV (range F))"
+  assumes F: "\<And>i. F i \<in> sets borel"
+  shows "borel = sigma UNIV (range F)"
+  using assms by (intro borel_eq_sigmaI1[where X="range (\<lambda>(i, j). G i j)" and F=F]) auto
+
+lemma borel_eq_sigmaI5:
+  fixes F :: "'i \<Rightarrow> 'j \<Rightarrow> 'a::topological_space set" and G :: "'l \<Rightarrow> 'a::topological_space set"
+  assumes borel_eq: "borel = sigma UNIV (range G)"
+  assumes X: "\<And>i. G i \<in> sets (sigma UNIV (range (\<lambda>(i, j). F i j)))"
+  assumes F: "\<And>i j. F i j \<in> sets borel"
+  shows "borel = sigma UNIV (range (\<lambda>(i, j). F i j))"
+  using assms by (intro borel_eq_sigmaI1[where X="range G" and F="(\<lambda>(i, j). F i j)"]) auto
 
 lemma 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")
+  "{x\<Colon>'a. a < x $$ i} \<in> sigma_sets UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. x $$ i < a}))"
+  (is "?set \<in> ?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)})"
+  interpret sigma_algebra UNIV ?SIGMA
+    by (intro sigma_algebra_sigma_sets) simp_all
+  have *: "?set = (\<Union>n. UNIV - {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"]
@@ -381,100 +384,78 @@
     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> {S. open S}"
-    then have "open S" unfolding mem_Collect_eq .
-    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
+  show "?set \<in> ?SIGMA" unfolding *
+    by (auto intro!: Diff)
 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 borel_eq_halfspace_less:
+  "borel = sigma UNIV (range (\<lambda>(a, i). {x::'a::ordered_euclidean_space. x $$ i < a}))"
+  (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI3[OF borel_def])
+  fix S :: "'a set" assume "S \<in> {S. open S}"
+  then have "open S" by simp
+  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> ?SIGMA"
+    unfolding *
+    by (safe intro!: countable_UN Int countable_INT) (auto intro!: halfspace_gt_in_halfspace)
+qed auto
 
-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 borel_eq_halfspace_le:
+  "borel = sigma UNIV (range (\<lambda> (a, i). {x::'a::ordered_euclidean_space. x $$ i \<le> a}))"
+  (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
+  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> ?SIGMA" unfolding *
+    by (safe intro!: countable_UN) auto
+qed auto
 
-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 borel_eq_halfspace_ge:
+  "borel = sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a \<le> x $$ i}))"
+  (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_less])
+  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> ?SIGMA" unfolding *
+      by (safe intro!: compl_sets) auto
+qed auto
 
-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"
+lemma borel_eq_halfspace_greater:
+  "borel = sigma UNIV (range (\<lambda> (a, i). {x\<Colon>'a\<Colon>ordered_euclidean_space. a < x $$ i}))"
+  (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI2[OF borel_eq_halfspace_le])
+  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> ?SIGMA" unfolding *
+    by (safe intro!: compl_sets) auto
+qed auto
+
+lemma borel_eq_atMost:
+  "borel = sigma UNIV (range (\<lambda>a. {..a\<Colon>'a\<Colon>ordered_euclidean_space}))"
+  (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
+  fix a i show "{x. x$$i \<le> a} \<in> ?SIGMA"
   proof cases
-    fix a i assume "i < DIM('a)"
+    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
@@ -484,28 +465,19 @@
       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
+    show "{x. x$$i \<le> a} \<in> ?SIGMA" unfolding *
+      by (safe intro!: countable_UN) auto
+  qed (auto intro: sigma_sets_top sigma_sets.Empty)
+qed auto
 
-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"
+lemma borel_eq_greaterThan:
+  "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {a<..}))"
+  (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_le])
+  fix a i show "{x. x$$i \<le> a} \<in> ?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
+    assume "i < DIM('a)"
+    have "{x::'a. x$$i \<le> a} = UNIV - {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
@@ -518,30 +490,22 @@
       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"
+    finally show "{x. x$$i \<le> a} \<in> ?SIGMA"
       apply (simp only:)
       apply (safe intro!: countable_UN Diff)
-      apply (auto simp: sets_sigma intro!: sigma_sets.Basic)
+      apply (auto intro: sigma_sets_top)
       done
-  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
+  qed (auto intro: sigma_sets_top sigma_sets.Empty)
+qed auto
 
-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"
+lemma borel_eq_lessThan:
+  "borel = sigma UNIV (range (\<lambda>a\<Colon>'a\<Colon>ordered_euclidean_space. {..<a}))"
+  (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI4[OF borel_eq_halfspace_ge])
+  fix a i show "{x. a \<le> x$$i} \<in> ?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
+    have "{x::'a. a \<le> x$$i} = UNIV - {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
@@ -554,184 +518,70 @@
       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"
+    finally show "{x. a \<le> x$$i} \<in> ?SIGMA"
       apply (simp only:)
       apply (safe intro!: countable_UN Diff)
-      apply (auto simp: sets_sigma intro!: sigma_sets.Basic)
+      apply (auto intro: sigma_sets_top)
       done
-  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 borel_eq_atMost:
-  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> a. {.. a::'a\<Colon>ordered_euclidean_space})\<rparr>)"
-    (is "_ = ?SIGMA")
-proof (intro algebra.equality 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 intro: sigma_sets_top sigma_sets.Empty)
 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 (intro algebra.equality 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 (intro algebra.equality 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 (intro algebra.equality 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
+  "borel = sigma UNIV (range (\<lambda>(a,b). {a..b} \<Colon>'a\<Colon>ordered_euclidean_space set))"
+  (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI5[OF borel_eq_atMost])
+  fix a::'a
+  have *: "{..a} = (\<Union>n::nat. {- real n *\<^sub>R One .. a})"
+  proof (safe, simp_all add: eucl_le[where 'a='a])
+    fix x
+    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> ?SIGMA" unfolding *
+    by (safe intro!: countable_UN)
+       (auto intro!: sigma_sets_top)
 qed auto
 
 lemma borel_eq_greaterThanLessThan:
-  "borel = (sigma \<lparr>space=UNIV, sets=range (\<lambda> (a, b). {a <..< (b :: 'a \<Colon> ordered_euclidean_space)})\<rparr>)"
+  "borel = sigma UNIV (range (\<lambda> (a, b). {a <..< b} :: 'a \<Colon> ordered_euclidean_space set))"
     (is "_ = ?SIGMA")
-proof (intro algebra.equality 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> {S. open S}"
-      then have "open M" by simp
-      have "M \<in> sets ?SIGMA"
-        apply (subst open_UNION[OF `open M`])
-        apply (safe intro!: countable_UN)
-        apply (auto simp add: sigma_def intro!: sigma_sets.Basic)
-        done }
-    then show ?thesis
-      unfolding borel_def by (intro sets_sigma_subset) auto
-  qed
+proof (rule borel_eq_sigmaI1[OF borel_def])
+  fix M :: "'a set" assume "M \<in> {S. open S}"
+  then have "open M" by simp
+  show "M \<in> ?SIGMA"
+    apply (subst open_UNION[OF `open M`])
+    apply (safe intro!: countable_UN)
+    apply auto
+    done
 qed auto
 
 lemma borel_eq_atLeastLessThan:
-  "borel = sigma \<lparr>space=UNIV, sets=range (\<lambda>(a, b). {a ..< b :: real})\<rparr>" (is "_ = ?S")
-proof (intro algebra.equality antisym)
-  interpret sigma_algebra ?S
-    by (rule sigma_algebra_sigma) auto
-  show "sets borel \<subseteq> sets ?S"
-    unfolding borel_eq_lessThan
-  proof (intro sets_sigma_subset subsetI)
-    have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto
-    fix A :: "real set" assume "A \<in> sets \<lparr>space = UNIV, sets = range lessThan\<rparr>"
-    then obtain x where "A = {..< x}" by auto
-    then have "A = (\<Union>i::nat. {-real i ..< x})"
-      by (auto simp: move_uminus real_arch_simple)
-    then show "A \<in> sets ?S"
-      by (auto simp: sets_sigma intro!: sigma_sets.intros)
-  qed simp
-  show "sets ?S \<subseteq> sets borel"
-    by (intro borel.sets_sigma_subset) auto
-qed simp_all
-
-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 (intro algebra.equality 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
+  "borel = sigma UNIV (range (\<lambda>(a, b). {a ..< b :: real}))" (is "_ = ?SIGMA")
+proof (rule borel_eq_sigmaI5[OF borel_eq_lessThan])
+  have move_uminus: "\<And>x y::real. -x \<le> y \<longleftrightarrow> -y \<le> x" by auto
+  fix x :: real
+  have "{..<x} = (\<Union>i::nat. {-real i ..< x})"
+    by (auto simp: move_uminus real_arch_simple)
+  then show "{..< x} \<in> ?SIGMA"
+    by (auto intro: sigma_sets.intros)
 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 (intro algebra.equality 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 (intro algebra.equality 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 (intro algebra.equality 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:
+lemma 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"
+  assumes F: "borel = sigma UNIV (range F)"
+  and S_eq: "\<And>a i. S a i = f -` F (a,i) \<inter> space M" 
+  and S: "\<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)
+    by (auto intro!: measurable_sets sigma_sets.Basic simp: assms(1))
 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"
@@ -740,61 +590,58 @@
       with a show ?thesis unfolding assms(2) by simp
     next
       assume "\<not> i < DIM('c)"
-      from assms(3)[OF this] show ?thesis .
+      from S[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
+  then show "f \<in> borel_measurable M"
+    by (auto intro!: measurable_measure_of simp: S_eq F)
 qed
 
-lemma (in sigma_algebra) borel_measurable_iff_halfspace_le:
+lemma 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:
+lemma 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:
+lemma 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:
+lemma 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
+  by (rule borel_measurable_halfspacesI[OF borel_eq_halfspace_greater]) auto
 
-lemma (in sigma_algebra) borel_measurable_iff_le:
+lemma borel_measurable_iff_le:
   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w \<le> a} \<in> sets M)"
   using borel_measurable_iff_halfspace_le[where 'c=real] by simp
 
-lemma (in sigma_algebra) borel_measurable_iff_less:
+lemma borel_measurable_iff_less:
   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. f w < a} \<in> sets M)"
   using borel_measurable_iff_halfspace_less[where 'c=real] by simp
 
-lemma (in sigma_algebra) borel_measurable_iff_ge:
+lemma borel_measurable_iff_ge:
   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a \<le> f w} \<in> sets M)"
   using borel_measurable_iff_halfspace_ge[where 'c=real] by simp
 
-lemma (in sigma_algebra) borel_measurable_iff_greater:
+lemma borel_measurable_iff_greater:
   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
   using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
 
 lemma borel_measurable_euclidean_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)
+proof (rule borel_measurableI)
   fix S::"real set" assume "open S"
   from open_vimage_euclidean_component[OF this]
-  show "(\<lambda>x. x $$ i) -` S \<in> sets borel"
+  show "(\<lambda>x. x $$ i) -` S \<inter> space borel \<in> sets borel"
     by (auto intro: borel_open)
 qed
 
-lemma (in sigma_algebra) borel_measurable_euclidean_space:
+lemma borel_measurable_euclidean_space:
   fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M)"
 proof safe
@@ -810,7 +657,7 @@
 
 subsection "Borel measurable operators"
 
-lemma (in sigma_algebra) affine_borel_measurable_vector:
+lemma affine_borel_measurable_vector:
   fixes f :: "'a \<Rightarrow> 'x::real_normed_vector"
   assumes "f \<in> borel_measurable M"
   shows "(\<lambda>x. a + b *\<^sub>R f x) \<in> borel_measurable M"
@@ -821,8 +668,7 @@
     assume "b \<noteq> 0"
     with `open S` have "open ((\<lambda>x. (- a + x) /\<^sub>R b) ` S)" (is "open ?S")
       by (auto intro!: open_affinity simp: scaleR_add_right)
-    hence "?S \<in> sets borel"
-      unfolding borel_def by (auto simp: sigma_def intro!: sigma_sets.Basic)
+    hence "?S \<in> sets borel" by auto
     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)
@@ -831,13 +677,13 @@
   qed simp
 qed
 
-lemma (in sigma_algebra) affine_borel_measurable:
+lemma 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]:
+lemma borel_measurable_add[simp, intro]:
   fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -855,7 +701,7 @@
     by (simp add: borel_measurable_iff_ge)
 qed
 
-lemma (in sigma_algebra) borel_measurable_setsum[simp, intro]:
+lemma borel_measurable_setsum[simp, intro]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
@@ -864,7 +710,7 @@
   thus ?thesis using assms by induct auto
 qed simp
 
-lemma (in sigma_algebra) borel_measurable_square:
+lemma borel_measurable_square:
   fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   shows "(\<lambda>x. (f x)^2) \<in> borel_measurable M"
@@ -916,7 +762,7 @@
    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]:
+lemma borel_measurable_uminus[simp, intro]:
   fixes g :: "'a \<Rightarrow> real"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. - g x) \<in> borel_measurable M"
@@ -928,7 +774,7 @@
   finally show ?thesis .
 qed
 
-lemma (in sigma_algebra) borel_measurable_times[simp, intro]:
+lemma borel_measurable_times[simp, intro]:
   fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -947,7 +793,7 @@
     using 1 2 by simp
 qed
 
-lemma (in sigma_algebra) borel_measurable_setprod[simp, intro]:
+lemma borel_measurable_setprod[simp, intro]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> real"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
@@ -956,14 +802,14 @@
   thus ?thesis using assms by induct auto
 qed simp
 
-lemma (in sigma_algebra) borel_measurable_diff[simp, intro]:
+lemma borel_measurable_diff[simp, intro]:
   fixes f :: "'a \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
   unfolding diff_minus using assms by fast
 
-lemma (in sigma_algebra) borel_measurable_inverse[simp, intro]:
+lemma 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"
@@ -978,7 +824,7 @@
     by (auto intro!: Int Un)
 qed
 
-lemma (in sigma_algebra) borel_measurable_divide[simp, intro]:
+lemma borel_measurable_divide[simp, intro]:
   fixes f :: "'a \<Rightarrow> real"
   assumes "f \<in> borel_measurable M"
   and "g \<in> borel_measurable M"
@@ -986,7 +832,7 @@
   unfolding field_divide_inverse
   by (rule borel_measurable_inverse borel_measurable_times assms)+
 
-lemma (in sigma_algebra) borel_measurable_max[intro, simp]:
+lemma borel_measurable_max[intro, simp]:
   fixes f g :: "'a \<Rightarrow> real"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
@@ -1001,7 +847,7 @@
     by (auto intro!: Int)
 qed
 
-lemma (in sigma_algebra) borel_measurable_min[intro, simp]:
+lemma borel_measurable_min[intro, simp]:
   fixes f g :: "'a \<Rightarrow> real"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
@@ -1016,7 +862,7 @@
     by (auto intro!: Int)
 qed
 
-lemma (in sigma_algebra) borel_measurable_abs[simp, intro]:
+lemma borel_measurable_abs[simp, intro]:
   assumes "f \<in> borel_measurable M"
   shows "(\<lambda>x. \<bar>f x :: real\<bar>) \<in> borel_measurable M"
 proof -
@@ -1033,14 +879,14 @@
   fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
   assumes "continuous_on UNIV f"
   shows "f \<in> borel_measurable borel"
-  apply(rule borel.borel_measurableI)
+  apply(rule borel_measurableI)
   using continuous_open_preimage[OF assms] unfolding vimage_def by auto
 
 lemma borel_measurable_continuous_on:
   fixes f :: "'a::topological_space \<Rightarrow> 'b::t1_space"
   assumes cont: "continuous_on A f" "open A"
   shows "(\<lambda>x. if x \<in> A then f x else c) \<in> borel_measurable borel" (is "?f \<in> _")
-proof (rule borel.borel_measurableI)
+proof (rule borel_measurableI)
   fix S :: "'b set" assume "open S"
   then have "open {x\<in>A. f x \<in> S}"
     by (intro continuous_open_preimage[OF cont]) auto
@@ -1049,11 +895,11 @@
     {x\<in>A. f x \<in> S} \<union> (if c \<in> S then space borel - A else {})"
     by (auto split: split_if_asm)
   also have "\<dots> \<in> sets borel"
-    using * `open A` by (auto simp del: space_borel intro!: borel.Un)
+    using * `open A` by (auto simp del: space_borel intro!: Un)
   finally show "?f -` S \<inter> space borel \<in> sets borel" .
 qed
 
-lemma (in sigma_algebra) convex_measurable:
+lemma convex_measurable:
   fixes a b :: real
   assumes X: "X \<in> borel_measurable M" "X ` space M \<subseteq> { a <..< b}"
   assumes q: "convex_on { a <..< b} q"
@@ -1091,7 +937,7 @@
   finally show ?thesis .
 qed
 
-lemma (in sigma_algebra) borel_measurable_log[simp,intro]:
+lemma borel_measurable_log[simp,intro]:
   assumes f: "f \<in> borel_measurable M" and "1 < b"
   shows "(\<lambda>x. log b (f x)) \<in> borel_measurable M"
   using measurable_comp[OF f borel_measurable_borel_log[OF `1 < b`]]
@@ -1101,25 +947,21 @@
 
 lemma borel_measurable_ereal_borel:
   "ereal \<in> borel_measurable borel"
-  unfolding borel_def[where 'a=ereal]
-proof (rule borel.measurable_sigma)
-  fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = {S. open S} \<rparr>"
-  then have "open X" by simp
+proof (rule borel_measurableI)
+  fix X :: "ereal set" assume "open X"
   then have "open (ereal -` X \<inter> space borel)"
     by (simp add: open_ereal_vimage)
   then show "ereal -` X \<inter> space borel \<in> sets borel" by auto
-qed auto
+qed
 
-lemma (in sigma_algebra) borel_measurable_ereal[simp, intro]:
+lemma borel_measurable_ereal[simp, intro]:
   assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
   using measurable_comp[OF f borel_measurable_ereal_borel] unfolding comp_def .
 
 lemma borel_measurable_real_of_ereal_borel:
   "(real :: ereal \<Rightarrow> real) \<in> borel_measurable borel"
-  unfolding borel_def[where 'a=real]
-proof (rule borel.measurable_sigma)
-  fix B :: "real set" assume "B \<in> sets \<lparr>space = UNIV, sets = {S. open S} \<rparr>"
-  then have "open B" by simp
+proof (rule borel_measurableI)
+  fix B :: "real set" assume "open B"
   have *: "ereal -` real -` (B - {0}) = B - {0}" by auto
   have open_real: "open (real -` (B - {0}) :: ereal set)"
     unfolding open_ereal_def * using `open B` by auto
@@ -1137,13 +979,13 @@
     then show "(real -` B :: ereal set) \<inter> space borel \<in> sets borel"
       using open_real by auto
   qed
-qed auto
+qed
 
-lemma (in sigma_algebra) borel_measurable_real_of_ereal[simp, intro]:
+lemma borel_measurable_real_of_ereal[simp, intro]:
   assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. real (f x :: ereal)) \<in> borel_measurable M"
   using measurable_comp[OF f borel_measurable_real_of_ereal_borel] unfolding comp_def .
 
-lemma (in sigma_algebra) borel_measurable_ereal_iff:
+lemma borel_measurable_ereal_iff:
   shows "(\<lambda>x. ereal (f x)) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M"
 proof
   assume "(\<lambda>x. ereal (f x)) \<in> borel_measurable M"
@@ -1151,7 +993,7 @@
   show "f \<in> borel_measurable M" by auto
 qed auto
 
-lemma (in sigma_algebra) borel_measurable_ereal_iff_real:
+lemma borel_measurable_ereal_iff_real:
   fixes f :: "'a \<Rightarrow> ereal"
   shows "f \<in> borel_measurable M \<longleftrightarrow>
     ((\<lambda>x. real (f x)) \<in> borel_measurable M \<and> f -` {\<infinity>} \<inter> space M \<in> sets M \<and> f -` {-\<infinity>} \<inter> space M \<in> sets M)"
@@ -1165,7 +1007,7 @@
   finally show "f \<in> borel_measurable M" .
 qed (auto intro: measurable_sets borel_measurable_real_of_ereal)
 
-lemma (in sigma_algebra) less_eq_ge_measurable:
+lemma less_eq_ge_measurable:
   fixes f :: "'a \<Rightarrow> 'c::linorder"
   shows "f -` {a <..} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {..a} \<inter> space M \<in> sets M"
 proof
@@ -1178,7 +1020,7 @@
   ultimately show "f -` {a <..} \<inter> space M \<in> sets M" by auto
 qed
 
-lemma (in sigma_algebra) greater_eq_le_measurable:
+lemma greater_eq_le_measurable:
   fixes f :: "'a \<Rightarrow> 'c::linorder"
   shows "f -` {..< a} \<inter> space M \<in> sets M \<longleftrightarrow> f -` {a ..} \<inter> space M \<in> sets M"
 proof
@@ -1191,28 +1033,27 @@
   ultimately show "f -` {a ..} \<inter> space M \<in> sets M" by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_uminus_borel_ereal:
+lemma borel_measurable_uminus_borel_ereal:
   "(uminus :: ereal \<Rightarrow> ereal) \<in> borel_measurable borel"
-proof (subst borel_def, rule borel.measurable_sigma)
-  fix X :: "ereal set" assume "X \<in> sets \<lparr>space = UNIV, sets = {S. open S}\<rparr>"
-  then have "open X" by simp
+proof (rule borel_measurableI)
+  fix X :: "ereal set" assume "open X"
   have "uminus -` X = uminus ` X" by (force simp: image_iff)
   then have "open (uminus -` X)" using `open X` ereal_open_uminus by auto
   then show "uminus -` X \<inter> space borel \<in> sets borel" by auto
-qed auto
+qed
 
-lemma (in sigma_algebra) borel_measurable_uminus_ereal[intro]:
+lemma borel_measurable_uminus_ereal[intro]:
   assumes "f \<in> borel_measurable M"
   shows "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M"
   using measurable_comp[OF assms borel_measurable_uminus_borel_ereal] by (simp add: comp_def)
 
-lemma (in sigma_algebra) borel_measurable_uminus_eq_ereal[simp]:
+lemma borel_measurable_uminus_eq_ereal[simp]:
   "(\<lambda>x. - f x :: ereal) \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable M" (is "?l = ?r")
 proof
   assume ?l from borel_measurable_uminus_ereal[OF this] show ?r by simp
 qed auto
 
-lemma (in sigma_algebra) borel_measurable_eq_atMost_ereal:
+lemma borel_measurable_eq_atMost_ereal:
   fixes f :: "'a \<Rightarrow> ereal"
   shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..a} \<inter> space M \<in> sets M)"
 proof (intro iffI allI)
@@ -1244,7 +1085,7 @@
   qed
 qed (simp add: measurable_sets)
 
-lemma (in sigma_algebra) borel_measurable_eq_atLeast_ereal:
+lemma borel_measurable_eq_atLeast_ereal:
   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a..} \<inter> space M \<in> sets M)"
 proof
   assume pos: "\<forall>a. f -` {a..} \<inter> space M \<in> sets M"
@@ -1255,15 +1096,15 @@
   then show "f \<in> borel_measurable M" by simp
 qed (simp add: measurable_sets)
 
-lemma (in sigma_algebra) borel_measurable_ereal_iff_less:
+lemma borel_measurable_ereal_iff_less:
   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {..< a} \<inter> space M \<in> sets M)"
   unfolding borel_measurable_eq_atLeast_ereal greater_eq_le_measurable ..
 
-lemma (in sigma_algebra) borel_measurable_ereal_iff_ge:
+lemma borel_measurable_ereal_iff_ge:
   "(f::'a \<Rightarrow> ereal) \<in> borel_measurable M \<longleftrightarrow> (\<forall>a. f -` {a <..} \<inter> space M \<in> sets M)"
   unfolding borel_measurable_eq_atMost_ereal less_eq_ge_measurable ..
 
-lemma (in sigma_algebra) borel_measurable_ereal_eq_const:
+lemma borel_measurable_ereal_eq_const:
   fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
   shows "{x\<in>space M. f x = c} \<in> sets M"
 proof -
@@ -1271,7 +1112,7 @@
   then show ?thesis using assms by (auto intro!: measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_ereal_neq_const:
+lemma borel_measurable_ereal_neq_const:
   fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
   shows "{x\<in>space M. f x \<noteq> c} \<in> sets M"
 proof -
@@ -1279,7 +1120,7 @@
   then show ?thesis using assms by (auto intro!: measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_ereal_le[intro,simp]:
+lemma borel_measurable_ereal_le[intro,simp]:
   fixes f g :: "'a \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -1294,7 +1135,7 @@
   with f g show ?thesis by (auto intro!: Un simp: measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_ereal_less[intro,simp]:
+lemma borel_measurable_ereal_less[intro,simp]:
   fixes f :: "'a \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -1304,7 +1145,7 @@
   then show ?thesis using g f by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_ereal_eq[intro,simp]:
+lemma borel_measurable_ereal_eq[intro,simp]:
   fixes f :: "'a \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -1314,7 +1155,7 @@
   then show ?thesis using g f by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_ereal_neq[intro,simp]:
+lemma borel_measurable_ereal_neq[intro,simp]:
   fixes f :: "'a \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable M"
   assumes g: "g \<in> borel_measurable M"
@@ -1324,12 +1165,12 @@
   thus ?thesis using f g by auto
 qed
 
-lemma (in sigma_algebra) split_sets:
+lemma split_sets:
   "{x\<in>space M. P x \<or> Q x} = {x\<in>space M. P x} \<union> {x\<in>space M. Q x}"
   "{x\<in>space M. P x \<and> Q x} = {x\<in>space M. P x} \<inter> {x\<in>space M. Q x}"
   by auto
 
-lemma (in sigma_algebra) borel_measurable_ereal_add[intro, simp]:
+lemma borel_measurable_ereal_add[intro, simp]:
   fixes f :: "'a \<Rightarrow> ereal"
   assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x + g x) \<in> borel_measurable M"
@@ -1344,7 +1185,7 @@
              intro!: Un measurable_If measurable_sets)
 qed
 
-lemma (in sigma_algebra) borel_measurable_ereal_setsum[simp, intro]:
+lemma borel_measurable_ereal_setsum[simp, intro]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Sum>i\<in>S. f i x) \<in> borel_measurable M"
@@ -1354,7 +1195,7 @@
     by induct auto
 qed (simp add: borel_measurable_const)
 
-lemma (in sigma_algebra) borel_measurable_ereal_abs[intro, simp]:
+lemma borel_measurable_ereal_abs[intro, simp]:
   fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M"
   shows "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M"
 proof -
@@ -1362,7 +1203,7 @@
   then show ?thesis using assms by (auto intro!: measurable_If)
 qed
 
-lemma (in sigma_algebra) borel_measurable_ereal_times[intro, simp]:
+lemma borel_measurable_ereal_times[intro, simp]:
   fixes f :: "'a \<Rightarrow> ereal" assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x * g x) \<in> borel_measurable M"
 proof -
@@ -1386,7 +1227,7 @@
        (auto simp: split_sets intro!: Int)
 qed
 
-lemma (in sigma_algebra) borel_measurable_ereal_setprod[simp, intro]:
+lemma borel_measurable_ereal_setprod[simp, intro]:
   fixes f :: "'c \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> S \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. \<Prod>i\<in>S. f i x) \<in> borel_measurable M"
@@ -1395,21 +1236,21 @@
   thus ?thesis using assms by induct auto
 qed simp
 
-lemma (in sigma_algebra) borel_measurable_ereal_min[simp, intro]:
+lemma borel_measurable_ereal_min[simp, intro]:
   fixes f g :: "'a \<Rightarrow> ereal"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
   shows "(\<lambda>x. min (g x) (f x)) \<in> borel_measurable M"
   using assms unfolding min_def by (auto intro!: measurable_If)
 
-lemma (in sigma_algebra) borel_measurable_ereal_max[simp, intro]:
+lemma borel_measurable_ereal_max[simp, intro]:
   fixes f g :: "'a \<Rightarrow> ereal"
   assumes "f \<in> borel_measurable M"
   and "g \<in> borel_measurable M"
   shows "(\<lambda>x. max (g x) (f x)) \<in> borel_measurable M"
   using assms unfolding max_def by (auto intro!: measurable_If)
 
-lemma (in sigma_algebra) borel_measurable_SUP[simp, intro]:
+lemma borel_measurable_SUP[simp, intro]:
   fixes f :: "'d\<Colon>countable \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. SUP i : A. f i x) \<in> borel_measurable M" (is "?sup \<in> borel_measurable M")
@@ -1422,7 +1263,7 @@
     using assms by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_INF[simp, intro]:
+lemma borel_measurable_INF[simp, intro]:
   fixes f :: "'d :: countable \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. i \<in> A \<Longrightarrow> f i \<in> borel_measurable M"
   shows "(\<lambda>x. INF i : A. f i x) \<in> borel_measurable M" (is "?inf \<in> borel_measurable M")
@@ -1435,26 +1276,39 @@
     using assms by auto
 qed
 
-lemma (in sigma_algebra) borel_measurable_liminf[simp, intro]:
+lemma borel_measurable_liminf[simp, intro]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. liminf (\<lambda>i. f i x)) \<in> borel_measurable M"
   unfolding liminf_SUPR_INFI using assms by auto
 
-lemma (in sigma_algebra) borel_measurable_limsup[simp, intro]:
+lemma borel_measurable_limsup[simp, intro]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. f i \<in> borel_measurable M"
   shows "(\<lambda>x. limsup (\<lambda>i. f i x)) \<in> borel_measurable M"
   unfolding limsup_INFI_SUPR using assms by auto
 
-lemma (in sigma_algebra) borel_measurable_ereal_diff[simp, intro]:
+lemma borel_measurable_ereal_diff[simp, intro]:
   fixes f g :: "'a \<Rightarrow> ereal"
   assumes "f \<in> borel_measurable M"
   assumes "g \<in> borel_measurable M"
   shows "(\<lambda>x. f x - g x) \<in> borel_measurable M"
   unfolding minus_ereal_def using assms by auto
 
-lemma (in sigma_algebra) borel_measurable_psuminf[simp, intro]:
+lemma borel_measurable_ereal_inverse[simp, intro]:
+  assumes f: "f \<in> borel_measurable M" shows "(\<lambda>x. inverse (f x) :: ereal) \<in> borel_measurable M"
+proof -
+  { fix x have "inverse (f x) = (if f x = 0 then \<infinity> else ereal (inverse (real (f x))))"
+      by (cases "f x") auto }
+  with f show ?thesis
+    by (auto intro!: measurable_If)
+qed
+
+lemma borel_measurable_ereal_divide[simp, intro]:
+  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. f x / g x :: ereal) \<in> borel_measurable M"
+  unfolding divide_ereal_def by auto
+
+lemma borel_measurable_psuminf[simp, intro]:
   fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
   assumes "\<And>i. f i \<in> borel_measurable M" and pos: "\<And>i x. x \<in> space M \<Longrightarrow> 0 \<le> f i x"
   shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
@@ -1465,7 +1319,7 @@
 
 section "LIMSEQ is borel measurable"
 
-lemma (in sigma_algebra) borel_measurable_LIMSEQ:
+lemma 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"
--- a/src/HOL/Probability/Caratheodory.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Caratheodory.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -6,7 +6,7 @@
 header {*Caratheodory Extension Theorem*}
 
 theory Caratheodory
-imports Sigma_Algebra "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits"
+  imports Measure_Space
 begin
 
 lemma sums_def2:
@@ -53,128 +53,53 @@
 
 subsection {* Measure Spaces *}
 
-record 'a measure_space = "'a algebra" +
-  measure :: "'a set \<Rightarrow> ereal"
-
-definition positive where "positive M f \<longleftrightarrow> f {} = (0::ereal) \<and> (\<forall>A\<in>sets M. 0 \<le> f A)"
-
-definition additive where "additive M f \<longleftrightarrow>
-  (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) = f x + f y)"
-
-definition countably_additive :: "('a, 'b) algebra_scheme \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
-  "countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
-    (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
-
-definition increasing where "increasing M f \<longleftrightarrow>
-  (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<subseteq> y \<longrightarrow> f x \<le> f y)"
-
 definition subadditive where "subadditive M f \<longleftrightarrow>
-  (\<forall>x \<in> sets M. \<forall>y \<in> sets M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
+  (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
 
 definition countably_subadditive where "countably_subadditive M f \<longleftrightarrow>
-  (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow>
+  (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow>
     (f (\<Union>i. A i) \<le> (\<Sum>i. f (A i))))"
 
-definition lambda_system where "lambda_system M f = {l \<in> sets M.
-  \<forall>x \<in> sets M. f (l \<inter> x) + f ((space M - l) \<inter> x) = f x}"
+definition lambda_system where "lambda_system \<Omega> M f = {l \<in> M.
+  \<forall>x \<in> M. f (l \<inter> x) + f ((\<Omega> - l) \<inter> x) = f x}"
 
 definition outer_measure_space where "outer_measure_space M f \<longleftrightarrow>
   positive M f \<and> increasing M f \<and> countably_subadditive M f"
 
 definition measure_set where "measure_set M f X = {r.
-  \<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
-
-locale measure_space = sigma_algebra M for M :: "('a, 'b) measure_space_scheme" +
-  assumes measure_positive: "positive M (measure M)"
-      and ca: "countably_additive M (measure M)"
-
-abbreviation (in measure_space) "\<mu> \<equiv> measure M"
-
-lemma (in measure_space)
-  shows empty_measure[simp, intro]: "\<mu> {} = 0"
-  and positive_measure[simp, intro!]: "\<And>A. A \<in> sets M \<Longrightarrow> 0 \<le> \<mu> A"
-  using measure_positive unfolding positive_def by auto
-
-lemma increasingD:
-  "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>sets M \<Longrightarrow> y\<in>sets M \<Longrightarrow> f x \<le> f y"
-  by (auto simp add: increasing_def)
+  \<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> X \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
 
 lemma subadditiveD:
-  "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> sets M \<Longrightarrow> y \<in> sets M
-    \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
+  "subadditive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) \<le> f x + f y"
   by (auto simp add: subadditive_def)
 
-lemma additiveD:
-  "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> sets M \<Longrightarrow> y \<in> sets M
-    \<Longrightarrow> f (x \<union> y) = f x + f y"
-  by (auto simp add: additive_def)
-
-lemma countably_additiveI:
-  assumes "\<And>A x. range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> sets M
-    \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
-  shows "countably_additive M f"
-  using assms by (simp add: countably_additive_def)
-
-section "Extend binary sets"
-
-lemma LIMSEQ_binaryset:
-  assumes f: "f {} = 0"
-  shows  "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) ----> f A + f B"
-proof -
-  have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
-    proof
-      fix n
-      show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
-        by (induct n)  (auto simp add: binaryset_def f)
-    qed
-  moreover
-  have "... ----> f A + f B" by (rule tendsto_const)
-  ultimately
-  have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
-    by metis
-  hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) ----> f A + f B"
-    by simp
-  thus ?thesis by (rule LIMSEQ_offset [where k=2])
-qed
-
-lemma binaryset_sums:
-  assumes f: "f {} = 0"
-  shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
-    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
-
-lemma suminf_binaryset_eq:
-  fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
-  shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
-  by (metis binaryset_sums sums_unique)
-
 subsection {* Lambda Systems *}
 
 lemma (in algebra) lambda_system_eq:
-  shows "lambda_system M f = {l \<in> sets M.
-    \<forall>x \<in> sets M. f (x \<inter> l) + f (x - l) = f x}"
+  shows "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (x \<inter> l) + f (x - l) = f x}"
 proof -
-  have [simp]: "!!l x. l \<in> sets M \<Longrightarrow> x \<in> sets M \<Longrightarrow> (space M - l) \<inter> x = x - l"
+  have [simp]: "!!l x. l \<in> M \<Longrightarrow> x \<in> M \<Longrightarrow> (\<Omega> - l) \<inter> x = x - l"
     by (metis Int_Diff Int_absorb1 Int_commute sets_into_space)
   show ?thesis
     by (auto simp add: lambda_system_def) (metis Int_commute)+
 qed
 
 lemma (in algebra) lambda_system_empty:
-  "positive M f \<Longrightarrow> {} \<in> lambda_system M f"
+  "positive M f \<Longrightarrow> {} \<in> lambda_system \<Omega> M f"
   by (auto simp add: positive_def lambda_system_eq)
 
 lemma lambda_system_sets:
-  "x \<in> lambda_system M f \<Longrightarrow> x \<in> sets M"
+  "x \<in> lambda_system \<Omega> M f \<Longrightarrow> x \<in> M"
   by (simp add: lambda_system_def)
 
 lemma (in algebra) lambda_system_Compl:
   fixes f:: "'a set \<Rightarrow> ereal"
-  assumes x: "x \<in> lambda_system M f"
-  shows "space M - x \<in> lambda_system M f"
+  assumes x: "x \<in> lambda_system \<Omega> M f"
+  shows "\<Omega> - x \<in> lambda_system \<Omega> M f"
 proof -
-  have "x \<subseteq> space M"
+  have "x \<subseteq> \<Omega>"
     by (metis sets_into_space lambda_system_sets x)
-  hence "space M - (space M - x) = x"
+  hence "\<Omega> - (\<Omega> - x) = x"
     by (metis double_diff equalityE)
   with x show ?thesis
     by (force simp add: lambda_system_def ac_simps)
@@ -182,16 +107,16 @@
 
 lemma (in algebra) lambda_system_Int:
   fixes f:: "'a set \<Rightarrow> ereal"
-  assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
-  shows "x \<inter> y \<in> lambda_system M f"
+  assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
+  shows "x \<inter> y \<in> lambda_system \<Omega> M f"
 proof -
   from xl yl show ?thesis
   proof (auto simp add: positive_def lambda_system_eq Int)
     fix u
-    assume x: "x \<in> sets M" and y: "y \<in> sets M" and u: "u \<in> sets M"
-       and fx: "\<forall>z\<in>sets M. f (z \<inter> x) + f (z - x) = f z"
-       and fy: "\<forall>z\<in>sets M. f (z \<inter> y) + f (z - y) = f z"
-    have "u - x \<inter> y \<in> sets M"
+    assume x: "x \<in> M" and y: "y \<in> M" and u: "u \<in> M"
+       and fx: "\<forall>z\<in>M. f (z \<inter> x) + f (z - x) = f z"
+       and fy: "\<forall>z\<in>M. f (z \<inter> y) + f (z - y) = f z"
+    have "u - x \<inter> y \<in> M"
       by (metis Diff Diff_Int Un u x y)
     moreover
     have "(u - (x \<inter> y)) \<inter> y = u \<inter> y - x" by blast
@@ -216,20 +141,20 @@
 
 lemma (in algebra) lambda_system_Un:
   fixes f:: "'a set \<Rightarrow> ereal"
-  assumes xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
-  shows "x \<union> y \<in> lambda_system M f"
+  assumes xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
+  shows "x \<union> y \<in> lambda_system \<Omega> M f"
 proof -
-  have "(space M - x) \<inter> (space M - y) \<in> sets M"
+  have "(\<Omega> - x) \<inter> (\<Omega> - y) \<in> M"
     by (metis Diff_Un Un compl_sets lambda_system_sets xl yl)
   moreover
-  have "x \<union> y = space M - ((space M - x) \<inter> (space M - y))"
+  have "x \<union> y = \<Omega> - ((\<Omega> - x) \<inter> (\<Omega> - y))"
     by auto (metis subsetD lambda_system_sets sets_into_space xl yl)+
   ultimately show ?thesis
     by (metis lambda_system_Compl lambda_system_Int xl yl)
 qed
 
 lemma (in algebra) lambda_system_algebra:
-  "positive M f \<Longrightarrow> algebra (M\<lparr>sets := lambda_system M f\<rparr>)"
+  "positive M f \<Longrightarrow> algebra \<Omega> (lambda_system \<Omega> M f)"
   apply (auto simp add: algebra_iff_Un)
   apply (metis lambda_system_sets set_mp sets_into_space)
   apply (metis lambda_system_empty)
@@ -238,117 +163,62 @@
   done
 
 lemma (in algebra) lambda_system_strong_additive:
-  assumes z: "z \<in> sets M" and disj: "x \<inter> y = {}"
-      and xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
+  assumes z: "z \<in> M" and disj: "x \<inter> y = {}"
+      and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
   shows "f (z \<inter> (x \<union> y)) = f (z \<inter> x) + f (z \<inter> y)"
 proof -
   have "z \<inter> x = (z \<inter> (x \<union> y)) \<inter> x" using disj by blast
   moreover
   have "z \<inter> y = (z \<inter> (x \<union> y)) - x" using disj by blast
   moreover
-  have "(z \<inter> (x \<union> y)) \<in> sets M"
+  have "(z \<inter> (x \<union> y)) \<in> M"
     by (metis Int Un lambda_system_sets xl yl z)
   ultimately show ?thesis using xl yl
     by (simp add: lambda_system_eq)
 qed
 
-lemma (in algebra) lambda_system_additive:
-     "additive (M (|sets := lambda_system M f|)) f"
+lemma (in algebra) lambda_system_additive: "additive (lambda_system \<Omega> M f) f"
 proof (auto simp add: additive_def)
   fix x and y
   assume disj: "x \<inter> y = {}"
-     and xl: "x \<in> lambda_system M f" and yl: "y \<in> lambda_system M f"
-  hence  "x \<in> sets M" "y \<in> sets M" by (blast intro: lambda_system_sets)+
+     and xl: "x \<in> lambda_system \<Omega> M f" and yl: "y \<in> lambda_system \<Omega> M f"
+  hence  "x \<in> M" "y \<in> M" by (blast intro: lambda_system_sets)+
   thus "f (x \<union> y) = f x + f y"
     using lambda_system_strong_additive [OF top disj xl yl]
     by (simp add: Un)
 qed
 
-lemma (in ring_of_sets) disjointed_additive:
-  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> sets M" "incseq A"
-  shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
-proof (induct n)
-  case (Suc n)
-  then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
-    by simp
-  also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
-    using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_incseq)
-  also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
-    using `incseq A` by (auto dest: incseq_SucD simp: disjointed_incseq)
-  finally show ?case .
-qed simp
-
 lemma (in ring_of_sets) countably_subadditive_subadditive:
   assumes f: "positive M f" and cs: "countably_subadditive M f"
   shows  "subadditive M f"
 proof (auto simp add: subadditive_def)
   fix x y
-  assume x: "x \<in> sets M" and y: "y \<in> sets M" and "x \<inter> y = {}"
+  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
   hence "disjoint_family (binaryset x y)"
     by (auto simp add: disjoint_family_on_def binaryset_def)
-  hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
-         (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
+  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
+         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
          f (\<Union>i. binaryset x y i) \<le> (\<Sum> n. f (binaryset x y n))"
     using cs by (auto simp add: countably_subadditive_def)
-  hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
+  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
          f (x \<union> y) \<le> (\<Sum> n. f (binaryset x y n))"
     by (simp add: range_binaryset_eq UN_binaryset_eq)
   thus "f (x \<union> y) \<le>  f x + f y" using f x y
     by (auto simp add: Un o_def suminf_binaryset_eq positive_def)
 qed
 
-lemma (in ring_of_sets) additive_sum:
-  fixes A:: "nat \<Rightarrow> 'a set"
-  assumes f: "positive M f" and ad: "additive M f" and "finite S"
-      and A: "range A \<subseteq> sets M"
-      and disj: "disjoint_family_on A S"
-  shows  "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
-using `finite S` disj proof induct
-  case empty show ?case using f by (simp add: positive_def)
-next
-  case (insert s S)
-  then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
-    by (auto simp add: disjoint_family_on_def neq_iff)
-  moreover
-  have "A s \<in> sets M" using A by blast
-  moreover have "(\<Union>i\<in>S. A i) \<in> sets M"
-    using A `finite S` by auto
-  moreover
-  ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
-    using ad UNION_in_sets A by (auto simp add: additive_def)
-  with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
-    by (auto simp add: additive_def subset_insertI)
-qed
-
-lemma (in algebra) increasing_additive_bound:
-  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ereal"
-  assumes f: "positive M f" and ad: "additive M f"
-      and inc: "increasing M f"
-      and A: "range A \<subseteq> sets M"
-      and disj: "disjoint_family A"
-  shows  "(\<Sum>i. f (A i)) \<le> f (space M)"
-proof (safe intro!: suminf_bound)
-  fix N
-  note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
-  have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
-    by (rule additive_sum [OF f ad _ A]) (auto simp: disj_N)
-  also have "... \<le> f (space M)" using space_closed A
-    by (intro increasingD[OF inc] finite_UN) auto
-  finally show "(\<Sum>i<N. f (A i)) \<le> f (space M)" by simp
-qed (insert f A, auto simp: positive_def)
-
 lemma lambda_system_increasing:
- "increasing M f \<Longrightarrow> increasing (M (|sets := lambda_system M f|)) f"
+ "increasing M f \<Longrightarrow> increasing (lambda_system \<Omega> M f) f"
   by (simp add: increasing_def lambda_system_def)
 
 lemma lambda_system_positive:
-  "positive M f \<Longrightarrow> positive (M (|sets := lambda_system M f|)) f"
+  "positive M f \<Longrightarrow> positive (lambda_system \<Omega> M f) f"
   by (simp add: positive_def lambda_system_def)
 
 lemma (in algebra) lambda_system_strong_sum:
   fixes A:: "nat \<Rightarrow> 'a set" and f :: "'a set \<Rightarrow> ereal"
-  assumes f: "positive M f" and a: "a \<in> sets M"
-      and A: "range A \<subseteq> lambda_system M f"
+  assumes f: "positive M f" and a: "a \<in> M"
+      and A: "range A \<subseteq> lambda_system \<Omega> M f"
       and disj: "disjoint_family A"
   shows  "(\<Sum>i = 0..<n. f (a \<inter>A i)) = f (a \<inter> (\<Union>i\<in>{0..<n}. A i))"
 proof (induct n)
@@ -357,11 +227,11 @@
   case (Suc n)
   have 2: "A n \<inter> UNION {0..<n} A = {}" using disj
     by (force simp add: disjoint_family_on_def neq_iff)
-  have 3: "A n \<in> lambda_system M f" using A
+  have 3: "A n \<in> lambda_system \<Omega> M f" using A
     by blast
-  interpret l: algebra "M\<lparr>sets := lambda_system M f\<rparr>"
+  interpret l: algebra \<Omega> "lambda_system \<Omega> M f"
     using f by (rule lambda_system_algebra)
-  have 4: "UNION {0..<n} A \<in> lambda_system M f"
+  have 4: "UNION {0..<n} A \<in> lambda_system \<Omega> M f"
     using A l.UNION_in_sets by simp
   from Suc.hyps show ?case
     by (simp add: atLeastLessThanSuc lambda_system_strong_additive [OF a 2 3 4])
@@ -369,23 +239,23 @@
 
 lemma (in sigma_algebra) lambda_system_caratheodory:
   assumes oms: "outer_measure_space M f"
-      and A: "range A \<subseteq> lambda_system M f"
+      and A: "range A \<subseteq> lambda_system \<Omega> M f"
       and disj: "disjoint_family A"
-  shows  "(\<Union>i. A i) \<in> lambda_system M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
+  shows  "(\<Union>i. A i) \<in> lambda_system \<Omega> M f \<and> (\<Sum>i. f (A i)) = f (\<Union>i. A i)"
 proof -
   have pos: "positive M f" and inc: "increasing M f"
    and csa: "countably_subadditive M f"
     by (metis oms outer_measure_space_def)+
   have sa: "subadditive M f"
     by (metis countably_subadditive_subadditive csa pos)
-  have A': "range A \<subseteq> sets (M(|sets := lambda_system M f|))" using A
-    by simp
-  interpret ls: algebra "M\<lparr>sets := lambda_system M f\<rparr>"
+  have A': "\<And>S. A`S \<subseteq> (lambda_system \<Omega> M f)" using A
+    by auto
+  interpret ls: algebra \<Omega> "lambda_system \<Omega> M f"
     using pos by (rule lambda_system_algebra)
-  have A'': "range A \<subseteq> sets M"
+  have A'': "range A \<subseteq> M"
      by (metis A image_subset_iff lambda_system_sets)
 
-  have U_in: "(\<Union>i. A i) \<in> sets M"
+  have U_in: "(\<Union>i. A i) \<in> M"
     by (metis A'' countable_UN)
   have U_eq: "f (\<Union>i. A i) = (\<Sum>i. f (A i))"
   proof (rule antisym)
@@ -396,22 +266,22 @@
     show "(\<Sum>i. f (A i)) \<le> f (\<Union>i. A i)"
       using ls.additive_sum [OF lambda_system_positive[OF pos] lambda_system_additive _ A' dis]
       using A''
-      by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] allI countable_UN)
+      by (intro suminf_bound[OF _ *]) (auto intro!: increasingD[OF inc] countable_UN)
   qed
   {
     fix a
-    assume a [iff]: "a \<in> sets M"
+    assume a [iff]: "a \<in> M"
     have "f (a \<inter> (\<Union>i. A i)) + f (a - (\<Union>i. A i)) = f a"
     proof -
       show ?thesis
       proof (rule antisym)
-        have "range (\<lambda>i. a \<inter> A i) \<subseteq> sets M" using A''
+        have "range (\<lambda>i. a \<inter> A i) \<subseteq> M" using A''
           by blast
         moreover
         have "disjoint_family (\<lambda>i. a \<inter> A i)" using disj
           by (auto simp add: disjoint_family_on_def)
         moreover
-        have "a \<inter> (\<Union>i. A i) \<in> sets M"
+        have "a \<inter> (\<Union>i. A i) \<in> M"
           by (metis Int U_in a)
         ultimately
         have "f (a \<inter> (\<Union>i. A i)) \<le> (\<Sum>i. f (a \<inter> A i))"
@@ -424,11 +294,11 @@
         have "(\<Sum>i. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
           proof (intro suminf_bound_add allI)
             fix n
-            have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
+            have UNION_in: "(\<Union>i\<in>{0..<n}. A i) \<in> M"
               by (metis A'' UNION_in_sets)
             have le_fa: "f (UNION {0..<n} A \<inter> a) \<le> f a" using A''
               by (blast intro: increasingD [OF inc] A'' UNION_in_sets)
-            have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system M f"
+            have ls: "(\<Union>i\<in>{0..<n}. A i) \<in> lambda_system \<Omega> M f"
               using ls.UNION_in_sets by (simp add: A)
             hence eq_fa: "f a = f (a \<inter> (\<Union>i\<in>{0..<n}. A i)) + f (a - (\<Union>i\<in>{0..<n}. A i))"
               by (simp add: lambda_system_eq UNION_in)
@@ -437,9 +307,9 @@
             thus "(\<Sum>i<n. f (a \<inter> A i)) + f (a - (\<Union>i. A i)) \<le> f a"
               by (simp add: lambda_system_strong_sum pos A disj eq_fa add_left_mono atLeast0LessThan[symmetric])
           next
-            have "\<And>i. a \<inter> A i \<in> sets M" using A'' by auto
+            have "\<And>i. a \<inter> A i \<in> M" using A'' by auto
             then show "\<And>i. 0 \<le> f (a \<inter> A i)" using pos[unfolded positive_def] by auto
-            have "\<And>i. a - (\<Union>i. A i) \<in> sets M" using A'' by auto
+            have "\<And>i. a - (\<Union>i. A i) \<in> M" using A'' by auto
             then have "\<And>i. 0 \<le> f (a - (\<Union>i. A i))" using pos[unfolded positive_def] by auto
             then show "f (a - (\<Union>i. A i)) \<noteq> -\<infinity>" by auto
           qed
@@ -460,66 +330,29 @@
 
 lemma (in sigma_algebra) caratheodory_lemma:
   assumes oms: "outer_measure_space M f"
-  shows "measure_space \<lparr> space = space M, sets = lambda_system M f, measure = f \<rparr>"
-    (is "measure_space ?M")
+  defines "L \<equiv> lambda_system \<Omega> M f"
+  shows "measure_space \<Omega> L f"
 proof -
   have pos: "positive M f"
     by (metis oms outer_measure_space_def)
-  have alg: "algebra ?M"
+  have alg: "algebra \<Omega> L"
     using lambda_system_algebra [of f, OF pos]
-    by (simp add: algebra_iff_Un)
+    by (simp add: algebra_iff_Un L_def)
   then
-  have "sigma_algebra ?M"
+  have "sigma_algebra \<Omega> L"
     using lambda_system_caratheodory [OF oms]
-    by (simp add: sigma_algebra_disjoint_iff)
+    by (simp add: sigma_algebra_disjoint_iff L_def)
   moreover
-  have "measure_space_axioms ?M"
+  have "countably_additive L f" "positive L f"
     using pos lambda_system_caratheodory [OF oms]
-    by (simp add: measure_space_axioms_def positive_def lambda_system_sets
-                  countably_additive_def o_def)
+    by (auto simp add: lambda_system_sets L_def countably_additive_def positive_def)
   ultimately
   show ?thesis
-    by (simp add: measure_space_def)
-qed
-
-lemma (in ring_of_sets) additive_increasing:
-  assumes posf: "positive M f" and addf: "additive M f"
-  shows "increasing M f"
-proof (auto simp add: increasing_def)
-  fix x y
-  assume xy: "x \<in> sets M" "y \<in> sets M" "x \<subseteq> y"
-  then have "y - x \<in> sets M" by auto
-  then have "0 \<le> f (y-x)" using posf[unfolded positive_def] by auto
-  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono) auto
-  also have "... = f (x \<union> (y-x))" using addf
-    by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
-  also have "... = f y"
-    by (metis Un_Diff_cancel Un_absorb1 xy(3))
-  finally show "f x \<le> f y" by simp
-qed
-
-lemma (in ring_of_sets) countably_additive_additive:
-  assumes posf: "positive M f" and ca: "countably_additive M f"
-  shows "additive M f"
-proof (auto simp add: additive_def)
-  fix x y
-  assume x: "x \<in> sets M" and y: "y \<in> sets M" and "x \<inter> y = {}"
-  hence "disjoint_family (binaryset x y)"
-    by (auto simp add: disjoint_family_on_def binaryset_def)
-  hence "range (binaryset x y) \<subseteq> sets M \<longrightarrow>
-         (\<Union>i. binaryset x y i) \<in> sets M \<longrightarrow>
-         f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
-    using ca
-    by (simp add: countably_additive_def)
-  hence "{x,y,{}} \<subseteq> sets M \<longrightarrow> x \<union> y \<in> sets M \<longrightarrow>
-         f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
-    by (simp add: range_binaryset_eq UN_binaryset_eq)
-  thus "f (x \<union> y) = f x + f y" using posf x y
-    by (auto simp add: Un suminf_binaryset_eq positive_def)
+    using pos by (simp add: measure_space_def)
 qed
 
 lemma inf_measure_nonempty:
-  assumes f: "positive M f" and b: "b \<in> sets M" and a: "a \<subseteq> b" "{} \<in> sets M"
+  assumes f: "positive M f" and b: "b \<in> M" and a: "a \<subseteq> b" "{} \<in> M"
   shows "f b \<in> measure_set M f a"
 proof -
   let ?A = "\<lambda>i::nat. (if i = 0 then b else {})"
@@ -534,14 +367,14 @@
 
 lemma (in ring_of_sets) inf_measure_agrees:
   assumes posf: "positive M f" and ca: "countably_additive M f"
-      and s: "s \<in> sets M"
+      and s: "s \<in> M"
   shows "Inf (measure_set M f s) = f s"
   unfolding Inf_ereal_def
 proof (safe intro!: Greatest_equality)
   fix z
   assume z: "z \<in> measure_set M f s"
   from this obtain A where
-    A: "range A \<subseteq> sets M" and disj: "disjoint_family A"
+    A: "range A \<subseteq> M" and disj: "disjoint_family A"
     and "s \<subseteq> (\<Union>x. A x)" and si: "(\<Sum>i. f (A i)) = z"
     by (auto simp add: measure_set_def comp_def)
   hence seq: "s = (\<Union>i. A i \<inter> s)" by blast
@@ -549,11 +382,11 @@
     by (metis additive_increasing ca countably_additive_additive posf)
   have sums: "(\<Sum>i. f (A i \<inter> s)) = f (\<Union>i. A i \<inter> s)"
     proof (rule ca[unfolded countably_additive_def, rule_format])
-      show "range (\<lambda>n. A n \<inter> s) \<subseteq> sets M" using A s
+      show "range (\<lambda>n. A n \<inter> s) \<subseteq> M" using A s
         by blast
       show "disjoint_family (\<lambda>n. A n \<inter> s)" using disj
         by (auto simp add: disjoint_family_on_def)
-      show "(\<Union>i. A i \<inter> s) \<in> sets M" using A s
+      show "(\<Union>i. A i \<inter> s) \<in> M" using A s
         by (metis UN_extend_simps(4) s seq)
     qed
   hence "f s = (\<Sum>i. f (A i \<inter> s))"
@@ -562,7 +395,7 @@
     proof (rule suminf_le_pos)
       fix n show "f (A n \<inter> s) \<le> f (A n)" using A s
         by (force intro: increasingD [OF inc])
-      fix N have "A N \<inter> s \<in> sets M"  using A s by auto
+      fix N have "A N \<inter> s \<in> M"  using A s by auto
       then show "0 \<le> f (A N \<inter> s)" using posf unfolding positive_def by auto
     qed
   also have "... = z" by (rule si)
@@ -578,7 +411,7 @@
   assumes posf: "positive M f" "r \<in> measure_set M f X"
   shows "0 \<le> r"
 proof -
-  obtain A where "range A \<subseteq> sets M" and r: "r = (\<Sum>i. f (A i))"
+  obtain A where "range A \<subseteq> M" and r: "r = (\<Sum>i. f (A i))"
     using `r \<in> measure_set M f X` unfolding measure_set_def by auto
   then show "0 \<le> r" using posf unfolding r positive_def
     by (intro suminf_0_le) auto
@@ -593,26 +426,25 @@
 qed
 
 lemma inf_measure_empty:
-  assumes posf: "positive M f" and "{} \<in> sets M"
+  assumes posf: "positive M f" and "{} \<in> M"
   shows "Inf (measure_set M f {}) = 0"
 proof (rule antisym)
   show "Inf (measure_set M f {}) \<le> 0"
-    by (metis complete_lattice_class.Inf_lower `{} \<in> sets M`
+    by (metis complete_lattice_class.Inf_lower `{} \<in> M`
               inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def])
 qed (rule inf_measure_pos[OF posf])
 
 lemma (in ring_of_sets) inf_measure_positive:
-  assumes p: "positive M f" and "{} \<in> sets M"
+  assumes p: "positive M f" and "{} \<in> M"
   shows "positive M (\<lambda>x. Inf (measure_set M f x))"
 proof (unfold positive_def, intro conjI ballI)
   show "Inf (measure_set M f {}) = 0" using inf_measure_empty[OF assms] by auto
-  fix A assume "A \<in> sets M"
+  fix A assume "A \<in> M"
 qed (rule inf_measure_pos[OF p])
 
 lemma (in ring_of_sets) inf_measure_increasing:
   assumes posf: "positive M f"
-  shows "increasing \<lparr> space = space M, sets = Pow (space M) \<rparr>
-                    (\<lambda>x. Inf (measure_set M f x))"
+  shows "increasing (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))"
 apply (clarsimp simp add: increasing_def)
 apply (rule complete_lattice_class.Inf_greatest)
 apply (rule complete_lattice_class.Inf_lower)
@@ -621,13 +453,13 @@
 
 lemma (in ring_of_sets) inf_measure_le:
   assumes posf: "positive M f" and inc: "increasing M f"
-      and x: "x \<in> {r . \<exists>A. range A \<subseteq> sets M \<and> s \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
+      and x: "x \<in> {r . \<exists>A. range A \<subseteq> M \<and> s \<subseteq> (\<Union>i. A i) \<and> (\<Sum>i. f (A i)) = r}"
   shows "Inf (measure_set M f s) \<le> x"
 proof -
-  obtain A where A: "range A \<subseteq> sets M" and ss: "s \<subseteq> (\<Union>i. A i)"
+  obtain A where A: "range A \<subseteq> M" and ss: "s \<subseteq> (\<Union>i. A i)"
              and xeq: "(\<Sum>i. f (A i)) = x"
     using x by auto
-  have dA: "range (disjointed A) \<subseteq> sets M"
+  have dA: "range (disjointed A) \<subseteq> M"
     by (metis A range_disjointed_sets)
   have "\<forall>n. f (disjointed A n) \<le> f (A n)"
     by (metis increasingD [OF inc] UNIV_I dA image_subset_iff disjointed_subset A comp_def)
@@ -648,8 +480,8 @@
 
 lemma (in ring_of_sets) inf_measure_close:
   fixes e :: ereal
-  assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (space M)" and "Inf (measure_set M f s) \<noteq> \<infinity>"
-  shows "\<exists>A. range A \<subseteq> sets M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and>
+  assumes posf: "positive M f" and e: "0 < e" and ss: "s \<subseteq> (\<Omega>)" and "Inf (measure_set M f s) \<noteq> \<infinity>"
+  shows "\<exists>A. range A \<subseteq> M \<and> disjoint_family A \<and> s \<subseteq> (\<Union>i. A i) \<and>
                (\<Sum>i. f (A i)) \<le> Inf (measure_set M f s) + e"
 proof -
   from `Inf (measure_set M f s) \<noteq> \<infinity>` have fin: "\<bar>Inf (measure_set M f s)\<bar> \<noteq> \<infinity>"
@@ -662,22 +494,21 @@
 
 lemma (in ring_of_sets) inf_measure_countably_subadditive:
   assumes posf: "positive M f" and inc: "increasing M f"
-  shows "countably_subadditive (| space = space M, sets = Pow (space M) |)
-                  (\<lambda>x. Inf (measure_set M f x))"
+  shows "countably_subadditive (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))"
 proof (simp add: countably_subadditive_def, safe)
   fix A :: "nat \<Rightarrow> 'a set"
   let ?outer = "\<lambda>B. Inf (measure_set M f B)"
-  assume A: "range A \<subseteq> Pow (space M)"
+  assume A: "range A \<subseteq> Pow (\<Omega>)"
      and disj: "disjoint_family A"
-     and sb: "(\<Union>i. A i) \<subseteq> space M"
+     and sb: "(\<Union>i. A i) \<subseteq> \<Omega>"
 
   { fix e :: ereal assume e: "0 < e" and "\<forall>i. ?outer (A i) \<noteq> \<infinity>"
-    hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> sets M \<and> disjoint_family (BB n) \<and>
+    hence "\<exists>BB. \<forall>n. range (BB n) \<subseteq> M \<and> disjoint_family (BB n) \<and>
         A n \<subseteq> (\<Union>i. BB n i) \<and> (\<Sum>i. f (BB n i)) \<le> ?outer (A n) + e * (1/2)^(Suc n)"
       apply (safe intro!: choice inf_measure_close [of f, OF posf])
       using e sb by (auto simp: ereal_zero_less_0_iff one_ereal_def)
     then obtain BB
-      where BB: "\<And>n. (range (BB n) \<subseteq> sets M)"
+      where BB: "\<And>n. (range (BB n) \<subseteq> M)"
       and disjBB: "\<And>n. disjoint_family (BB n)"
       and sbBB: "\<And>n. A n \<subseteq> (\<Union>i. BB n i)"
       and BBle: "\<And>n. (\<Sum>i. f (BB n i)) \<le> ?outer (A n) + e * (1/2)^(Suc n)"
@@ -697,7 +528,7 @@
       finally show ?thesis .
     qed
     def C \<equiv> "(split BB) o prod_decode"
-    have C: "!!n. C n \<in> sets M"
+    have C: "!!n. C n \<in> M"
       apply (rule_tac p="prod_decode n" in PairE)
       apply (simp add: C_def)
       apply (metis BB subsetD rangeI)
@@ -744,9 +575,8 @@
 qed
 
 lemma (in ring_of_sets) inf_measure_outer:
-  "\<lbrakk> positive M f ; increasing M f \<rbrakk>
-   \<Longrightarrow> outer_measure_space \<lparr> space = space M, sets = Pow (space M) \<rparr>
-                          (\<lambda>x. Inf (measure_set M f x))"
+  "\<lbrakk> positive M f ; increasing M f \<rbrakk> \<Longrightarrow>
+    outer_measure_space (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))"
   using inf_measure_pos[of M f]
   by (simp add: outer_measure_space_def inf_measure_empty
                 inf_measure_increasing inf_measure_countably_subadditive positive_def)
@@ -754,14 +584,13 @@
 lemma (in ring_of_sets) algebra_subset_lambda_system:
   assumes posf: "positive M f" and inc: "increasing M f"
       and add: "additive M f"
-  shows "sets M \<subseteq> lambda_system \<lparr> space = space M, sets = Pow (space M) \<rparr>
-                                (\<lambda>x. Inf (measure_set M f x))"
+  shows "M \<subseteq> lambda_system \<Omega> (Pow \<Omega>) (\<lambda>x. Inf (measure_set M f x))"
 proof (auto dest: sets_into_space
             simp add: algebra.lambda_system_eq [OF algebra_Pow])
   fix x s
-  assume x: "x \<in> sets M"
-     and s: "s \<subseteq> space M"
-  have [simp]: "!!x. x \<in> sets M \<Longrightarrow> s \<inter> (space M - x) = s-x" using s
+  assume x: "x \<in> M"
+     and s: "s \<subseteq> \<Omega>"
+  have [simp]: "!!x. x \<in> M \<Longrightarrow> s \<inter> (\<Omega> - x) = s-x" using s
     by blast
   have "Inf (measure_set M f (s\<inter>x)) + Inf (measure_set M f (s-x))
         \<le> Inf (measure_set M f s)"
@@ -774,7 +603,7 @@
     show ?thesis
     proof (rule complete_lattice_class.Inf_greatest)
       fix r assume "r \<in> measure_set M f s"
-      then obtain A where A: "disjoint_family A" "range A \<subseteq> sets M" "s \<subseteq> (\<Union>i. A i)"
+      then obtain A where A: "disjoint_family A" "range A \<subseteq> M" "s \<subseteq> (\<Union>i. A i)"
         and r: "r = (\<Sum>i. f (A i))" unfolding measure_set_def by auto
       have "Inf (measure_set M f (s \<inter> x)) \<le> (\<Sum>i. f (A i \<inter> x))"
         unfolding measure_set_def
@@ -822,34 +651,31 @@
 qed
 
 lemma measure_down:
-  "measure_space N \<Longrightarrow> sigma_algebra M \<Longrightarrow> sets M \<subseteq> sets N \<Longrightarrow> measure N = measure M \<Longrightarrow> measure_space M"
-  by (simp add: measure_space_def measure_space_axioms_def positive_def
-                countably_additive_def)
+  "measure_space \<Omega> N \<mu> \<Longrightarrow> sigma_algebra \<Omega> M \<Longrightarrow> M \<subseteq> N \<Longrightarrow> measure_space \<Omega> M \<mu>"
+  by (simp add: measure_space_def positive_def countably_additive_def)
      blast
 
 theorem (in ring_of_sets) caratheodory:
   assumes posf: "positive M f" and ca: "countably_additive M f"
-  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
-            measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
 proof -
   have inc: "increasing M f"
     by (metis additive_increasing ca countably_additive_additive posf)
   let ?infm = "(\<lambda>x. Inf (measure_set M f x))"
-  def ls \<equiv> "lambda_system (|space = space M, sets = Pow (space M)|) ?infm"
-  have mls: "measure_space \<lparr>space = space M, sets = ls, measure = ?infm\<rparr>"
+  def ls \<equiv> "lambda_system \<Omega> (Pow \<Omega>) ?infm"
+  have mls: "measure_space \<Omega> ls ?infm"
     using sigma_algebra.caratheodory_lemma
             [OF sigma_algebra_Pow  inf_measure_outer [OF posf inc]]
     by (simp add: ls_def)
-  hence sls: "sigma_algebra (|space = space M, sets = ls, measure = ?infm|)"
+  hence sls: "sigma_algebra \<Omega> ls"
     by (simp add: measure_space_def)
-  have "sets M \<subseteq> ls"
+  have "M \<subseteq> ls"
     by (simp add: ls_def)
        (metis ca posf inc countably_additive_additive algebra_subset_lambda_system)
-  hence sgs_sb: "sigma_sets (space M) (sets M) \<subseteq> ls"
-    using sigma_algebra.sigma_sets_subset [OF sls, of "sets M"]
+  hence sgs_sb: "sigma_sets (\<Omega>) (M) \<subseteq> ls"
+    using sigma_algebra.sigma_sets_subset [OF sls, of "M"]
     by simp
-  have "measure_space \<lparr> space = space M, sets = sets (sigma M), measure = ?infm \<rparr>"
-    unfolding sigma_def
+  have "measure_space \<Omega> (sigma_sets \<Omega> M) ?infm"
     by (rule measure_down [OF mls], rule sigma_algebra_sigma_sets)
        (simp_all add: sgs_sb space_closed)
   thus ?thesis using inf_measure_agrees [OF posf ca]
@@ -861,12 +687,12 @@
 lemma (in ring_of_sets) countably_additive_iff_continuous_from_below:
   assumes f: "positive M f" "additive M f"
   shows "countably_additive M f \<longleftrightarrow>
-    (\<forall>A. range A \<subseteq> sets M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i))"
+    (\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i))"
   unfolding countably_additive_def
 proof safe
-  assume count_sum: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow> UNION UNIV A \<in> sets M \<longrightarrow> (\<Sum>i. f (A i)) = f (UNION UNIV A)"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "incseq A" "(\<Union>i. A i) \<in> sets M"
-  then have dA: "range (disjointed A) \<subseteq> sets M" by (auto simp: range_disjointed_sets)
+  assume count_sum: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> UNION UNIV A \<in> M \<longrightarrow> (\<Sum>i. f (A i)) = f (UNION UNIV A)"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
+  then have dA: "range (disjointed A) \<subseteq> M" by (auto simp: range_disjointed_sets)
   with count_sum[THEN spec, of "disjointed A"] A(3)
   have f_UN: "(\<Sum>i. f (disjointed A i)) = f (\<Union>i. A i)"
     by (auto simp: UN_disjointed_eq disjoint_family_disjointed)
@@ -880,12 +706,12 @@
     using disjointed_additive[OF f A(1,2)] .
   ultimately show "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)" by simp
 next
-  assume cont: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> sets M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "disjoint_family A" "(\<Union>i. A i) \<in> sets M"
+  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> incseq A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "disjoint_family A" "(\<Union>i. A i) \<in> M"
   have *: "(\<Union>n. (\<Union>i\<le>n. A i)) = (\<Union>i. A i)" by auto
   have "(\<lambda>n. f (\<Union>i\<le>n. A i)) ----> f (\<Union>i. A i)"
   proof (unfold *[symmetric], intro cont[rule_format])
-    show "range (\<lambda>i. \<Union> i\<le>i. A i) \<subseteq> sets M" "(\<Union>i. \<Union> i\<le>i. A i) \<in> sets M"
+    show "range (\<lambda>i. \<Union> i\<le>i. A i) \<subseteq> M" "(\<Union>i. \<Union> i\<le>i. A i) \<in> M"
       using A * by auto
   qed (force intro!: incseq_SucI)
   moreover have "\<And>n. f (\<Union>i\<le>n. A i) = (\<Sum>i\<le>n. f (A i))"
@@ -901,18 +727,18 @@
 
 lemma (in ring_of_sets) continuous_from_above_iff_empty_continuous:
   assumes f: "positive M f" "additive M f"
-  shows "(\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> sets M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))
-     \<longleftrightarrow> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0)"
+  shows "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))
+     \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0)"
 proof safe
-  assume cont: "(\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> sets M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
+  assume cont: "(\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) \<in> M \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> f (\<Inter>i. A i))"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) = {}" "\<forall>i. f (A i) \<noteq> \<infinity>"
   with cont[THEN spec, of A] show "(\<lambda>i. f (A i)) ----> 0"
     using `positive M f`[unfolded positive_def] by auto
 next
-  assume cont: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
-  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets M" "decseq A" "(\<Inter>i. A i) \<in> sets M" "\<forall>i. f (A i) \<noteq> \<infinity>"
+  assume cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<forall>i. f (A i) \<noteq> \<infinity>) \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> M" "decseq A" "(\<Inter>i. A i) \<in> M" "\<forall>i. f (A i) \<noteq> \<infinity>"
 
-  have f_mono: "\<And>a b. a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b"
+  have f_mono: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<subseteq> b \<Longrightarrow> f a \<le> f b"
     using additive_increasing[OF f] unfolding increasing_def by simp
 
   have decseq_fA: "decseq (\<lambda>i. f (A i))"
@@ -932,7 +758,7 @@
   note f_fin = this
   have "(\<lambda>i. f (A i - (\<Inter>i. A i))) ----> 0"
   proof (intro cont[rule_format, OF _ decseq _ f_fin])
-    show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> sets M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
+    show "range (\<lambda>i. A i - (\<Inter>i. A i)) \<subseteq> M" "(\<Inter>i. A i - (\<Inter>i. A i)) = {}"
       using A by auto
   qed
   from INF_Lim_ereal[OF decseq_f this]
@@ -956,17 +782,17 @@
 qed
 
 lemma positiveD1: "positive M f \<Longrightarrow> f {} = 0" by (auto simp: positive_def)
-lemma positiveD2: "positive M f \<Longrightarrow> A \<in> sets M \<Longrightarrow> 0 \<le> f A" by (auto simp: positive_def)
+lemma positiveD2: "positive M f \<Longrightarrow> A \<in> M \<Longrightarrow> 0 \<le> f A" by (auto simp: positive_def)
 
 lemma (in ring_of_sets) empty_continuous_imp_continuous_from_below:
-  assumes f: "positive M f" "additive M f" "\<forall>A\<in>sets M. f A \<noteq> \<infinity>"
-  assumes cont: "\<forall>A. range A \<subseteq> sets M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
-  assumes A: "range A \<subseteq> sets M" "incseq A" "(\<Union>i. A i) \<in> sets M"
+  assumes f: "positive M f" "additive M f" "\<forall>A\<in>M. f A \<noteq> \<infinity>"
+  assumes cont: "\<forall>A. range A \<subseteq> M \<longrightarrow> decseq A \<longrightarrow> (\<Inter>i. A i) = {} \<longrightarrow> (\<lambda>i. f (A i)) ----> 0"
+  assumes A: "range A \<subseteq> M" "incseq A" "(\<Union>i. A i) \<in> M"
   shows "(\<lambda>i. f (A i)) ----> f (\<Union>i. A i)"
 proof -
-  have "\<forall>A\<in>sets M. \<exists>x. f A = ereal x"
+  have "\<forall>A\<in>M. \<exists>x. f A = ereal x"
   proof
-    fix A assume "A \<in> sets M" with f show "\<exists>x. f A = ereal x"
+    fix A assume "A \<in> M" with f show "\<exists>x. f A = ereal x"
       unfolding positive_def by (cases "f A") auto
   qed
   from bchoice[OF this] guess f' .. note f' = this[rule_format]
@@ -991,20 +817,19 @@
 qed
 
 lemma (in ring_of_sets) empty_continuous_imp_countably_additive:
-  assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>sets M. f A \<noteq> \<infinity>"
-  assumes cont: "\<And>A. range A \<subseteq> sets M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
+  assumes f: "positive M f" "additive M f" and fin: "\<forall>A\<in>M. f A \<noteq> \<infinity>"
+  assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
   shows "countably_additive M f"
   using countably_additive_iff_continuous_from_below[OF f]
   using empty_continuous_imp_continuous_from_below[OF f fin] cont
   by blast
 
 lemma (in ring_of_sets) caratheodory_empty_continuous:
-  assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> sets M \<Longrightarrow> f A \<noteq> \<infinity>"
-  assumes cont: "\<And>A. range A \<subseteq> sets M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
-  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> sets M. \<mu> s = f s) \<and>
-            measure_space \<lparr> space = space M, sets = sets (sigma M), measure = \<mu> \<rparr>"
+  assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> M \<Longrightarrow> f A \<noteq> \<infinity>"
+  assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
+  shows "\<exists>\<mu> :: 'a set \<Rightarrow> ereal. (\<forall>s \<in> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
 proof (intro caratheodory empty_continuous_imp_countably_additive f)
-  show "\<forall>A\<in>sets M. f A \<noteq> \<infinity>" using fin by auto
+  show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
 qed (rule cont)
 
 end
--- a/src/HOL/Probability/Complete_Measure.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Complete_Measure.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -6,218 +6,231 @@
 imports Lebesgue_Integration
 begin
 
-locale completeable_measure_space = measure_space
-
-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
+  "split_completion M A p = (if A \<in> sets M then p = (A, {}) else
+   \<exists>N'. A = fst p \<union> snd p \<and> fst p \<inter> snd p = {} \<and> fst p \<in> sets M \<and> snd p \<subseteq> N' \<and> N' \<in> null_sets M)"
 
-definition (in completeable_measure_space)
-  "null_part A = snd (Eps (split_completion A))"
-
-abbreviation (in completeable_measure_space) "\<mu>' A \<equiv> \<mu> (main_part A)"
+definition
+  "main_part M A = fst (Eps (split_completion M A))"
 
-definition (in completeable_measure_space) completion :: "('a, 'b) measure_space_scheme" 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' },
-                  measure = \<mu>',
-                  \<dots> = more M \<rparr>"
-
+definition
+  "null_part M A = snd (Eps (split_completion M A))"
 
-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
+definition completion :: "'a measure \<Rightarrow> 'a measure" where
+  "completion M = measure_of (space M) { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }
+    (emeasure M \<circ> main_part M)"
 
-lemma (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 completion_into_space:
+  "{ S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' } \<subseteq> Pow (space M)"
+  using sets_into_space 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 space_completion[simp]: "space (completion M) = space M"
+  unfolding completion_def using space_measure_of[OF completion_into_space] by simp
 
-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'])
+lemma completionI:
+  assumes "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M"
+  shows "A \<in> { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }"
   using assms by auto
 
-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)
+lemma completionE:
+  assumes "A \<in> { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }"
+  obtains S N N' where "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M"
+  using assms by auto
+
+lemma sigma_algebra_completion:
+  "sigma_algebra (space M) { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }"
+    (is "sigma_algebra _ ?A")
+  unfolding sigma_algebra_iff2
+proof (intro conjI ballI allI impI)
+  show "?A \<subseteq> Pow (space M)"
+    using sets_into_space by auto
 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
+  show "{} \<in> ?A" by 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)
+  let ?C = "space M"
+  fix A assume "A \<in> ?A" from completionE[OF this] guess S N N' .
+  then show "space M - A \<in> ?A"
+    by (intro completionI[of _ "(?C - S) \<inter> (?C - N')" "(?C - S) \<inter> N' \<inter> (?C - N)"]) auto
+next
+  fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> ?A"
+  then have "\<forall>n. \<exists>S N N'. A n = S \<union> N \<and> S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N'"
+    by (auto simp: image_subset_iff)
   from choice[OF this] guess S ..
   from choice[OF this] guess N ..
   from choice[OF this] guess N' ..
-  then show "UNION UNIV A \<in> sets completion"
+  then show "UNION UNIV A \<in> ?A"
     using null_sets_UN[of N']
-    by (intro sets_completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"])
-       auto
-qed auto
+    by (intro completionI[of _ "UNION UNIV S" "UNION UNIV N" "UNION UNIV N'"]) auto
+qed
+
+lemma sets_completion:
+  "sets (completion M) = { S \<union> N |S N N'. S \<in> sets M \<and> N' \<in> null_sets M \<and> N \<subseteq> N' }"
+  using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_completion] by (simp add: completion_def)
+
+lemma sets_completionE:
+  assumes "A \<in> sets (completion M)"
+  obtains S N N' where "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M"
+  using assms unfolding sets_completion by auto
+
+lemma sets_completionI:
+  assumes "A = S \<union> N" "N \<subseteq> N'" "N' \<in> null_sets M" "S \<in> sets M"
+  shows "A \<in> sets (completion M)"
+  using assms unfolding sets_completion by auto
+
+lemma sets_completionI_sets[intro, simp]:
+  "A \<in> sets M \<Longrightarrow> A \<in> sets (completion M)"
+  unfolding sets_completion by force
 
-lemma (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
+lemma null_sets_completion:
+  assumes "N' \<in> null_sets M" "N \<subseteq> N'" shows "N \<in> sets (completion M)"
+  using assms by (intro sets_completionI[of N "{}" N N']) auto
+
+lemma split_completion:
+  assumes "A \<in> sets (completion M)"
+  shows "split_completion M A (main_part M A, null_part M A)"
+proof cases
+  assume "A \<in> sets M" then show ?thesis
+    by (simp add: split_completion_def[abs_def] main_part_def null_part_def)
+next
+  assume nA: "A \<notin> sets M"
+  show ?thesis
+    unfolding main_part_def null_part_def if_not_P[OF nA]
+  proof (rule someI2_ex)
+    from assms[THEN sets_completionE] guess S N N' . note A = this
+    let ?P = "(S, N - S)"
+    show "\<exists>p. split_completion M A p"
+      unfolding split_completion_def if_not_P[OF nA] using A
+    proof (intro exI conjI)
+      show "A = fst ?P \<union> snd ?P" using A by auto
+      show "snd ?P \<subseteq> N'" using A by auto
+   qed auto
   qed auto
-qed auto
+qed
 
-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
+  assumes "S \<in> sets (completion M)"
+  shows main_part_sets[intro, simp]: "main_part M S \<in> sets M"
+    and main_part_null_part_Un[simp]: "main_part M S \<union> null_part M S = S"
+    and main_part_null_part_Int[simp]: "main_part M S \<inter> null_part M S = {}"
+  using split_completion[OF assms]
+  by (auto simp: split_completion_def split: split_if_asm)
 
-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 main_part[simp]: "S \<in> sets M \<Longrightarrow> main_part M S = S"
+  using split_completion[of S M]
+  by (auto simp: split_completion_def split: split_if_asm)
 
-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"
+lemma null_part:
+  assumes "S \<in> sets (completion M)" shows "\<exists>N. N\<in>null_sets M \<and> null_part M S \<subseteq> N"
+  using split_completion[OF assms] by (auto simp: split_completion_def split: split_if_asm)
+
+lemma null_part_sets[intro, simp]:
+  assumes "S \<in> sets M" shows "null_part M S \<in> sets M" "emeasure M (null_part M S) = 0"
 proof -
-  have S: "S \<in> sets completion" using assms by auto
-  have "S - main_part S \<in> sets M" using assms by auto
+  have S: "S \<in> sets (completion M)" using assms by auto
+  have "S - main_part M S \<in> sets M" using assms by auto
   moreover
   from main_part_null_part_Un[OF S] main_part_null_part_Int[OF S]
-  have "S - main_part S = null_part S" by auto
-  ultimately show sets: "null_part S \<in> sets M" by auto
+  have "S - main_part M S = null_part M S" by auto
+  ultimately show sets: "null_part M S \<in> sets M" by auto
   from null_part[OF S] guess N ..
-  with measure_eq_0[of N "null_part S"] sets
-  show "\<mu> (null_part S) = 0" by auto
-qed
-
-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>' S"
-    using S assms measure_additive[of "main_part S" "null_part S"]
-    by (auto simp: measure_additive)
-  finally show ?thesis by simp
+  with emeasure_eq_0[of N _ "null_part M S"] sets
+  show "emeasure M (null_part M S) = 0" by auto
 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:
+lemma emeasure_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))"
+  assumes "range S \<subseteq> sets (completion M)"
+  shows "emeasure M (main_part M (\<Union>i. (S i))) = emeasure M (\<Union>i. main_part M (S i))"
 proof -
-  have S: "\<And>i. S i \<in> sets completion" 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"
+  have S: "\<And>i. S i \<in> sets (completion M)" using assms by auto
+  then have UN: "(\<Union>i. S i) \<in> sets (completion M)" by auto
+  have "\<forall>i. \<exists>N. N \<in> null_sets M \<and> null_part M (S i) \<subseteq> N"
     using null_part[OF S] by auto
   from choice[OF this] guess N .. note N = this
-  then have UN_N: "(\<Union>i. N i) \<in> null_sets" by (intro null_sets_UN) auto
-  have "(\<Union>i. S i) \<in> sets completion" using S by auto
+  then have UN_N: "(\<Union>i. N i) \<in> null_sets M" by (intro null_sets_UN) auto
+  have "(\<Union>i. S i) \<in> sets (completion M)" using S by auto
   from null_part[OF this] guess N' .. note N' = this
   let ?N = "(\<Union>i. N i) \<union> N'"
-  have null_set: "?N \<in> null_sets" using N' UN_N by (intro nullsets.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"
+  have null_set: "?N \<in> null_sets M" using N' UN_N by (intro null_sets.Un) auto
+  have "main_part M (\<Union>i. S i) \<union> ?N = (main_part M (\<Union>i. S i) \<union> null_part M (\<Union>i. S i)) \<union> ?N"
     using N' by auto
-  also have "\<dots> = (\<Union>i. main_part (S i) \<union> null_part (S i)) \<union> ?N"
+  also have "\<dots> = (\<Union>i. main_part M (S i) \<union> null_part M (S i)) \<union> ?N"
     unfolding main_part_null_part_Un[OF S] main_part_null_part_Un[OF UN] by auto
-  also have "\<dots> = (\<Union>i. main_part (S i)) \<union> ?N"
+  also have "\<dots> = (\<Union>i. main_part M (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)"
+  finally have *: "main_part M (\<Union>i. S i) \<union> ?N = (\<Union>i. main_part M (S i)) \<union> ?N" .
+  have "emeasure M (main_part M (\<Union>i. S i)) = emeasure M (main_part M (\<Union>i. S i) \<union> ?N)"
+    using null_set UN by (intro emeasure_Un_null_set[symmetric]) auto
+  also have "\<dots> = emeasure M ((\<Union>i. main_part M (S i)) \<union> ?N)"
     unfolding * ..
-  also have "\<dots> = \<mu> (\<Union>i. main_part (S i))"
-    using null_set S by (intro measure_Un_null_set) auto
+  also have "\<dots> = emeasure M (\<Union>i. main_part M (S i))"
+    using null_set S by (intro emeasure_Un_null_set) auto
   finally show ?thesis .
 qed
 
-lemma (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
+lemma emeasure_completion[simp]:
+  assumes S: "S \<in> sets (completion M)" shows "emeasure (completion M) S = emeasure M (main_part M S)"
+proof (subst emeasure_measure_of[OF completion_def completion_into_space])
+  let ?\<mu> = "emeasure M \<circ> main_part M"
+  show "S \<in> sets (completion M)" "?\<mu> S = emeasure M (main_part M S) " using S by simp_all
+  show "positive (sets (completion M)) ?\<mu>"
+    by (simp add: positive_def emeasure_nonneg)
+  show "countably_additive (sets (completion M)) ?\<mu>"
+  proof (intro countably_additiveI)
+    fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets (completion M)" "disjoint_family A"
+    have "disjoint_family (\<lambda>i. main_part M (A i))"
+    proof (intro disjoint_family_on_bisimulation[OF A(2)])
+      fix n m assume "A n \<inter> A m = {}"
+      then have "(main_part M (A n) \<union> null_part M (A n)) \<inter> (main_part M (A m) \<union> null_part M (A m)) = {}"
+        using A by (subst (1 2) main_part_null_part_Un) auto
+      then show "main_part M (A n) \<inter> main_part M (A m) = {}" by auto
+    qed
+    then have "(\<Sum>n. emeasure M (main_part M (A n))) = emeasure M (\<Union>i. main_part M (A i))"
+      using A by (auto intro!: suminf_emeasure)
+    then show "(\<Sum>n. ?\<mu> (A n)) = ?\<mu> (UNION UNIV A)"
+      by (simp add: completion_def emeasure_main_part_UN[OF A(1)])
+  qed
 qed
 
-sublocale completeable_measure_space \<subseteq> completion!: measure_space completion
-  where "measure completion = \<mu>'"
-proof -
-  show "measure_space completion"
-  proof
-    show "positive completion (measure completion)"
-      by (auto simp: completion_def positive_def)
-  next
-    show "countably_additive completion (measure completion)"
-    proof (intro countably_additiveI)
-      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>n. measure completion (A n)) = \<mu> (\<Union>i. main_part (A i))"
-        unfolding completion_def using A by (auto intro!: measure_countably_additive)
-      then show "(\<Sum>n. measure completion (A n)) = measure completion (UNION UNIV A)"
-        by (simp add: completion_def \<mu>_main_part_UN[OF A(1)])
-    qed
-  qed
-  show "measure completion = \<mu>'" unfolding completion_def by simp
-qed
+lemma emeasure_completion_UN:
+  "range S \<subseteq> sets (completion M) \<Longrightarrow>
+    emeasure (completion M) (\<Union>i::nat. (S i)) = emeasure M (\<Union>i. main_part M (S i))"
+  by (subst emeasure_completion) (auto simp add: emeasure_main_part_UN)
 
-lemma (in completeable_measure_space) completion_ex_simple_function:
-  assumes f: "simple_function completion f"
-  shows "\<exists>f'. simple_function M f' \<and> (AE x. f x = f' x)"
+lemma emeasure_completion_Un:
+  assumes S: "S \<in> sets (completion M)" and T: "T \<in> sets (completion M)"
+  shows "emeasure (completion M) (S \<union> T) = emeasure M (main_part M S \<union> main_part M T)"
+proof (subst emeasure_completion)
+  have UN: "(\<Union>i. binary (main_part M S) (main_part M T) i) = (\<Union>i. main_part M (binary S T i))"
+    unfolding binary_def by (auto split: split_if_asm)
+  show "emeasure M (main_part M (S \<union> T)) = emeasure M (main_part M S \<union> main_part M T)"
+    using emeasure_main_part_UN[of "binary S T" M] assms
+    unfolding range_binary_eq Un_range_binary UN by auto
+qed (auto intro: S T)
+
+lemma sets_completionI_sub:
+  assumes N: "N' \<in> null_sets M" "N \<subseteq> N'"
+  shows "N \<in> sets (completion M)"
+  using assms by (intro sets_completionI[of _ "{}" N N']) auto
+
+lemma completion_ex_simple_function:
+  assumes f: "simple_function (completion M) f"
+  shows "\<exists>f'. simple_function M f' \<and> (AE x in M. f x = f' x)"
 proof -
   let ?F = "\<lambda>x. f -` {x} \<inter> space M"
-  have F: "\<And>x. ?F x \<in> sets completion" 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"
+  have F: "\<And>x. ?F x \<in> sets (completion M)" and fin: "finite (f`space M)"
+    using simple_functionD[OF f] simple_functionD[OF f] by simp_all
+  have "\<forall>x. \<exists>N. N \<in> null_sets M \<and> null_part M (?F x) \<subseteq> N"
     using F null_part by auto
   from choice[OF this] obtain N where
-    N: "\<And>x. null_part (?F x) \<subseteq> N x" "\<And>x. N x \<in> null_sets" by auto
+    N: "\<And>x. null_part M (?F x) \<subseteq> N x" "\<And>x. N x \<in> null_sets M" by auto
   let ?N = "\<Union>x\<in>f`space M. N x"
   let ?f' = "\<lambda>x. if x \<in> ?N then undefined else f x"
-  have sets: "?N \<in> null_sets" using N fin by (intro nullsets.finite_UN) auto
+  have sets: "?N \<in> null_sets M" using N fin by (intro null_sets.finite_UN) auto
   show ?thesis unfolding simple_function_def
   proof (safe intro!: exI[of _ ?f'])
     have "?f' ` space M \<subseteq> f`space M \<union> {undefined}" by auto
-    from finite_subset[OF this] completion.simple_functionD(1)[OF f]
+    from finite_subset[OF this] simple_functionD(1)[OF f]
     show "finite (?f' ` space M)" by auto
   next
     fix x assume "x \<in> space M"
@@ -225,13 +238,13 @@
       (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)
+      using N(2) sets_into_space by (auto split: split_if_asm simp: null_sets_def)
     moreover { fix y have "?F y \<union> ?N \<in> sets M"
       proof cases
         assume y: "y \<in> f`space M"
-        have "?F y \<union> ?N = (main_part (?F y) \<union> null_part (?F y)) \<union> ?N"
+        have "?F y \<union> ?N = (main_part M (?F y) \<union> null_part M (?F y)) \<union> ?N"
           using main_part_null_part_Un[OF F] by auto
-        also have "\<dots> = main_part (?F y) \<union> ?N"
+        also have "\<dots> = main_part M (?F y) \<union> ?N"
           using y N by auto
         finally show ?thesis
           using F sets by auto
@@ -240,34 +253,34 @@
         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"
+      have "?F (f x) - ?N = main_part M (?F (f x)) \<union> null_part M (?F (f x)) - ?N"
         using main_part_null_part_Un[OF F] by auto
-      also have "\<dots> = main_part (?F (f x)) - ?N"
+      also have "\<dots> = main_part M (?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"
+    show "AE x in M. f x = ?f' x"
       by (rule AE_I', rule sets) auto
   qed
 qed
 
-lemma (in completeable_measure_space) completion_ex_borel_measurable_pos:
+lemma completion_ex_borel_measurable_pos:
   fixes g :: "'a \<Rightarrow> ereal"
-  assumes g: "g \<in> borel_measurable completion" and "\<And>x. 0 \<le> g x"
-  shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
+  assumes g: "g \<in> borel_measurable (completion M)" and "\<And>x. 0 \<le> g x"
+  shows "\<exists>g'\<in>borel_measurable M. (AE x in M. g x = g' x)"
 proof -
-  from g[THEN completion.borel_measurable_implies_simple_function_sequence'] guess f . note f = this
+  from g[THEN borel_measurable_implies_simple_function_sequence'] guess f . note f = this
   from this(1)[THEN completion_ex_simple_function]
-  have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x. f i x = f' x)" ..
+  have "\<forall>i. \<exists>f'. simple_function M f' \<and> (AE x in M. f i x = f' x)" ..
   from this[THEN choice] obtain f' where
     sf: "\<And>i. simple_function M (f' i)" and
-    AE: "\<forall>i. AE x. f i x = f' i x" by auto
+    AE: "\<forall>i. AE x in M. f i x = f' i x" by auto
   show ?thesis
   proof (intro bexI)
     from AE[unfolded AE_all_countable[symmetric]]
-    show "AE x. g x = (SUP i. f' i x)" (is "AE x. g x = ?f x")
+    show "AE x in M. g x = (SUP i. f' i x)" (is "AE x in M. g x = ?f x")
     proof (elim AE_mp, safe intro!: AE_I2)
       fix x assume eq: "\<forall>i. f i x = f' i x"
       moreover have "g x = (SUP i. f i x)"
@@ -279,20 +292,20 @@
   qed
 qed
 
-lemma (in completeable_measure_space) completion_ex_borel_measurable:
+lemma completion_ex_borel_measurable:
   fixes g :: "'a \<Rightarrow> ereal"
-  assumes g: "g \<in> borel_measurable completion"
-  shows "\<exists>g'\<in>borel_measurable M. (AE x. g x = g' x)"
+  assumes g: "g \<in> borel_measurable (completion M)"
+  shows "\<exists>g'\<in>borel_measurable M. (AE x in M. g x = g' x)"
 proof -
-  have "(\<lambda>x. max 0 (g x)) \<in> borel_measurable completion" "\<And>x. 0 \<le> max 0 (g x)" using g by auto
+  have "(\<lambda>x. max 0 (g x)) \<in> borel_measurable (completion M)" "\<And>x. 0 \<le> max 0 (g x)" using g by auto
   from completion_ex_borel_measurable_pos[OF this] guess g_pos ..
   moreover
-  have "(\<lambda>x. max 0 (- g x)) \<in> borel_measurable completion" "\<And>x. 0 \<le> max 0 (- g x)" using g by auto
+  have "(\<lambda>x. max 0 (- g x)) \<in> borel_measurable (completion M)" "\<And>x. 0 \<le> max 0 (- g x)" using g by auto
   from completion_ex_borel_measurable_pos[OF this] guess g_neg ..
   ultimately
   show ?thesis
   proof (safe intro!: bexI[of _ "\<lambda>x. g_pos x - g_neg x"])
-    show "AE x. max 0 (- g x) = g_neg x \<longrightarrow> max 0 (g x) = g_pos x \<longrightarrow> g x = g_pos x - g_neg x"
+    show "AE x in M. max 0 (- g x) = g_neg x \<longrightarrow> max 0 (g x) = g_pos x \<longrightarrow> g x = g_pos x - g_neg x"
     proof (intro AE_I2 impI)
       fix x assume g: "max 0 (- g x) = g_neg x" "max 0 (g x) = g_pos x"
       show "g x = g_pos x - g_neg x" unfolding g[symmetric]
--- a/src/HOL/Probability/Conditional_Probability.thy	Mon Apr 23 12:23:23 2012 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,161 +0,0 @@
-(*  Title:      HOL/Probability/Conditional_Probability.thy
-    Author:     Johannes Hölzl, TU München
-*)
-
-header {*Conditional probability*}
-
-theory Conditional_Probability
-imports Probability_Measure Radon_Nikodym
-begin
-
-section "Conditional Expectation and Probability"
-
-definition (in prob_space)
-  "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable N \<and> (\<forall>x. 0 \<le> Y x)
-    \<and> (\<forall>C\<in>sets N. (\<integral>\<^isup>+x. Y x * indicator C x\<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x\<partial>M)))"
-
-lemma (in prob_space) conditional_expectation_exists:
-  fixes X :: "'a \<Rightarrow> ereal" and N :: "('a, 'b) measure_space_scheme"
-  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
-  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
-  shows "\<exists>Y\<in>borel_measurable N. (\<forall>x. 0 \<le> Y x) \<and> (\<forall>C\<in>sets N.
-      (\<integral>\<^isup>+x. Y x * indicator C x \<partial>M) = (\<integral>\<^isup>+x. X x * indicator C x \<partial>M))"
-proof -
-  note N(4)[simp]
-  interpret P: prob_space N
-    using prob_space_subalgebra[OF N] .
-
-  let ?f = "\<lambda>A x. X x * indicator A x"
-  let ?Q = "\<lambda>A. integral\<^isup>P M (?f A)"
-
-  from measure_space_density[OF borel]
-  have Q: "measure_space (N\<lparr> measure := ?Q \<rparr>)"
-    apply (rule measure_space.measure_space_subalgebra[of "M\<lparr> measure := ?Q \<rparr>"])
-    using N by (auto intro!: P.sigma_algebra_cong)
-  then interpret Q: measure_space "N\<lparr> measure := ?Q \<rparr>" .
-
-  have "P.absolutely_continuous ?Q"
-    unfolding P.absolutely_continuous_def
-  proof safe
-    fix A assume "A \<in> sets N" "P.\<mu> A = 0"
-    then have f_borel: "?f A \<in> borel_measurable M" "AE x. x \<notin> A"
-      using borel N by (auto intro!: borel_measurable_indicator AE_not_in)
-    then show "?Q A = 0"
-      by (auto simp add: positive_integral_0_iff_AE)
-  qed
-  from P.Radon_Nikodym[OF Q this]
-  obtain Y where Y: "Y \<in> borel_measurable N" "\<And>x. 0 \<le> Y x"
-    "\<And>A. A \<in> sets N \<Longrightarrow> ?Q A =(\<integral>\<^isup>+x. Y x * indicator A x \<partial>N)"
-    by blast
-  with N(2) show ?thesis
-    by (auto intro!: bexI[OF _ Y(1)] simp: positive_integral_subalgebra[OF _ _ N(2,3,4,1)])
-qed
-
-lemma (in prob_space)
-  fixes X :: "'a \<Rightarrow> ereal" and N :: "('a, 'b) measure_space_scheme"
-  assumes borel: "X \<in> borel_measurable M" "AE x. 0 \<le> X x"
-  and N: "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M" "\<And>A. measure N A = \<mu> A"
-  shows borel_measurable_conditional_expectation:
-    "conditional_expectation N X \<in> borel_measurable N"
-  and conditional_expectation: "\<And>C. C \<in> sets N \<Longrightarrow>
-      (\<integral>\<^isup>+x. conditional_expectation N X x * indicator C x \<partial>M) =
-      (\<integral>\<^isup>+x. X x * indicator C x \<partial>M)"
-   (is "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C")
-proof -
-  note CE = conditional_expectation_exists[OF assms, unfolded Bex_def]
-  then show "conditional_expectation N X \<in> borel_measurable N"
-    unfolding conditional_expectation_def by (rule someI2_ex) blast
-
-  from CE show "\<And>C. C \<in> sets N \<Longrightarrow> ?eq C"
-    unfolding conditional_expectation_def by (rule someI2_ex) blast
-qed
-
-lemma (in sigma_algebra) factorize_measurable_function_pos:
-  fixes Z :: "'a \<Rightarrow> ereal" and Y :: "'a \<Rightarrow> 'c"
-  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
-  assumes Z: "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)"
-  shows "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. max 0 (Z x) = g (Y x)"
-proof -
-  interpret M': sigma_algebra M' by fact
-  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
-  from M'.sigma_algebra_vimage[OF this]
-  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
-
-  from va.borel_measurable_implies_simple_function_sequence'[OF Z] guess f . note f = this
-
-  have "\<forall>i. \<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
-  proof
-    fix i
-    from f(1)[of i] have "finite (f i`space M)" and B_ex:
-      "\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> space M"
-      unfolding simple_function_def by auto
-    from B_ex[THEN bchoice] guess B .. note B = this
-
-    let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x"
-
-    show "\<exists>g. simple_function M' g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
-    proof (intro exI[of _ ?g] conjI ballI)
-      show "simple_function M' ?g" using B by auto
-
-      fix x assume "x \<in> space M"
-      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::ereal)"
-        unfolding indicator_def using B by auto
-      then show "f i x = ?g (Y x)" using `x \<in> space M` f(1)[of i]
-        by (subst va.simple_function_indicator_representation) auto
-    qed
-  qed
-  from choice[OF this] guess g .. note g = this
-
-  show ?thesis
-  proof (intro ballI bexI)
-    show "(\<lambda>x. SUP i. g i x) \<in> borel_measurable M'"
-      using g by (auto intro: M'.borel_measurable_simple_function)
-    fix x assume "x \<in> space M"
-    have "max 0 (Z x) = (SUP i. f i x)" using f by simp
-    also have "\<dots> = (SUP i. g i (Y x))"
-      using g `x \<in> space M` by simp
-    finally show "max 0 (Z x) = (SUP i. g i (Y x))" .
-  qed
-qed
-
-lemma (in sigma_algebra) factorize_measurable_function:
-  fixes Z :: "'a \<Rightarrow> ereal" and Y :: "'a \<Rightarrow> 'c"
-  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
-  shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
-    \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
-proof safe
-  interpret M': sigma_algebra M' by fact
-  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
-  from M'.sigma_algebra_vimage[OF this]
-  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
-
-  { fix g :: "'c \<Rightarrow> ereal" assume "g \<in> borel_measurable M'"
-    with M'.measurable_vimage_algebra[OF Y]
-    have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-      by (rule measurable_comp)
-    moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
-    then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
-       g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-       by (auto intro!: measurable_cong)
-    ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-      by simp }
-
-  assume Z: "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-  with assms have "(\<lambda>x. - Z x) \<in> borel_measurable M"
-    "(\<lambda>x. - Z x) \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
-    by auto
-  from factorize_measurable_function_pos[OF assms(1,2) this] guess n .. note n = this
-  from factorize_measurable_function_pos[OF assms Z] guess p .. note p = this
-  let ?g = "\<lambda>x. p x - n x"
-  show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
-  proof (intro bexI ballI)
-    show "?g \<in> borel_measurable M'" using p n by auto
-    fix x assume "x \<in> space M"
-    then have "p (Y x) = max 0 (Z x)" "n (Y x) = max 0 (- Z x)"
-      using p n by auto
-    then show "Z x = ?g (Y x)"
-      by (auto split: split_max)
-  qed
-qed
-
-end
\ No newline at end of file
--- a/src/HOL/Probability/Finite_Product_Measure.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Finite_Product_Measure.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -8,6 +8,9 @@
 imports Binary_Product_Measure
 begin
 
+lemma split_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
+  by auto
+
 lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)"
   unfolding Pi_def by auto
 
@@ -34,9 +37,6 @@
 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)"
@@ -86,7 +86,7 @@
   "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)
+  by (auto simp: restrict_def)
 
 lemma extensional_insert_undefined[intro, simp]:
   assumes "a \<in> extensional (insert i I)"
@@ -130,16 +130,16 @@
   using assms by (auto simp: restrict_Pi_cancel)
 
 lemma restrict_fupd[simp]: "i \<notin> I \<Longrightarrow> restrict (f (i := x)) I = restrict f I"
-  by (auto simp: restrict_def intro!: ext)
+  by (auto simp: restrict_def)
 
 lemma merge_restrict[simp]:
   "merge I (restrict x I) J y = merge I x J y"
   "merge I x J (restrict y J) = merge I x J y"
-  unfolding merge_def by (auto intro!: ext)
+  unfolding merge_def by auto
 
 lemma merge_x_x_eq_restrict[simp]:
   "merge I x J x = restrict x (I \<union> J)"
-  unfolding merge_def by (auto intro!: ext)
+  unfolding merge_def by auto
 
 lemma Pi_fupd_iff: "i \<in> I \<Longrightarrow> f \<in> Pi I (B(i := A)) \<longleftrightarrow> f \<in> Pi (I - {i}) B \<and> f i \<in> A"
   apply auto
@@ -233,339 +233,355 @@
 
 section "Products"
 
-locale product_sigma_algebra =
-  fixes M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme"
-  assumes sigma_algebras: "\<And>i. sigma_algebra (M i)"
+definition prod_emb where
+  "prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (PIE i:I. space (M i))"
+
+lemma prod_emb_iff: 
+  "f \<in> prod_emb I M K X \<longleftrightarrow> f \<in> extensional I \<and> (restrict f K \<in> X) \<and> (\<forall>i\<in>I. f i \<in> space (M i))"
+  unfolding prod_emb_def by auto
 
-locale finite_product_sigma_algebra = product_sigma_algebra M
-  for M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme" +
-  fixes I :: "'i set"
-  assumes finite_index[simp, intro]: "finite I"
+lemma
+  shows prod_emb_empty[simp]: "prod_emb M L K {} = {}"
+    and prod_emb_Un[simp]: "prod_emb M L K (A \<union> B) = prod_emb M L K A \<union> prod_emb M L K B"
+    and prod_emb_Int: "prod_emb M L K (A \<inter> B) = prod_emb M L K A \<inter> prod_emb M L K B"
+    and prod_emb_UN[simp]: "prod_emb M L K (\<Union>i\<in>I. F i) = (\<Union>i\<in>I. prod_emb M L K (F i))"
+    and prod_emb_INT[simp]: "I \<noteq> {} \<Longrightarrow> prod_emb M L K (\<Inter>i\<in>I. F i) = (\<Inter>i\<in>I. prod_emb M L K (F i))"
+    and prod_emb_Diff[simp]: "prod_emb M L K (A - B) = prod_emb M L K A - prod_emb M L K B"
+  by (auto simp: prod_emb_def)
 
-definition
-  "product_algebra_generator I M = \<lparr> space = (\<Pi>\<^isub>E i \<in> I. space (M i)),
-    sets = Pi\<^isub>E I ` (\<Pi> i \<in> I. sets (M i)),
-    measure = \<lambda>A. (\<Prod>i\<in>I. measure (M i) ((SOME F. A = Pi\<^isub>E I F) i)) \<rparr>"
+lemma prod_emb_PiE: "J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow>
+    prod_emb I M J (\<Pi>\<^isub>E i\<in>J. E i) = (\<Pi>\<^isub>E i\<in>I. if i \<in> J then E i else space (M i))"
+  by (force simp: prod_emb_def Pi_iff split_if_mem2)
+
+lemma prod_emb_PiE_same_index[simp]: "(\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> space (M i)) \<Longrightarrow> prod_emb I M I (Pi\<^isub>E I E) = Pi\<^isub>E I E"
+  by (auto simp: prod_emb_def Pi_iff)
 
-definition product_algebra_def:
-  "Pi\<^isub>M I M = sigma (product_algebra_generator I M)
-    \<lparr> measure := (SOME \<mu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<mu> \<rparr>) \<and>
-      (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))))\<rparr>"
+definition PiM :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
+  "PiM I M = extend_measure (\<Pi>\<^isub>E i\<in>I. space (M i))
+    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
+    (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^isub>E j\<in>J. X j))
+    (\<lambda>(J, X). \<Prod>j\<in>J \<union> {i\<in>I. emeasure (M i) (space (M i)) \<noteq> 1}. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
+
+definition prod_algebra :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i \<Rightarrow> 'a) set set" where
+  "prod_algebra I M = (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^isub>E j\<in>J. X j)) `
+    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}"
+
+abbreviation
+  "Pi\<^isub>M I M \<equiv> PiM I M"
 
 syntax
-  "_PiM"  :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
-              ('i => 'a, 'b) measure_space_scheme"  ("(3PIM _:_./ _)" 10)
+  "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3PIM _:_./ _)" 10)
 
 syntax (xsymbols)
-  "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
-             ('i => 'a, 'b) measure_space_scheme"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"   10)
+  "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"  10)
 
 syntax (HTML output)
-  "_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] =>
-             ('i => 'a, 'b) measure_space_scheme"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"   10)
+  "_PiM" :: "pttrn \<Rightarrow> 'i set \<Rightarrow> 'a measure \<Rightarrow> ('i => 'a) measure"  ("(3\<Pi>\<^isub>M _\<in>_./ _)"  10)
 
 translations
-  "PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)"
-
-abbreviation (in finite_product_sigma_algebra) "G \<equiv> product_algebra_generator I M"
-abbreviation (in finite_product_sigma_algebra) "P \<equiv> Pi\<^isub>M I M"
-
-sublocale product_sigma_algebra \<subseteq> M: sigma_algebra "M i" for i by (rule sigma_algebras)
-
-lemma sigma_into_space:
-  assumes "sets M \<subseteq> Pow (space M)"
-  shows "sets (sigma M) \<subseteq> Pow (space M)"
-  using sigma_sets_into_sp[OF assms] unfolding sigma_def by auto
+  "PIM x:I. M" == "CONST PiM I (%x. M)"
 
-lemma (in product_sigma_algebra) product_algebra_generator_into_space:
-  "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))"
-  using M.sets_into_space unfolding product_algebra_generator_def
-  by auto blast
-
-lemma (in product_sigma_algebra) product_algebra_into_space:
-  "sets (Pi\<^isub>M I M) \<subseteq> Pow (space (Pi\<^isub>M I M))"
-  using product_algebra_generator_into_space
-  by (auto intro!: sigma_into_space simp add: product_algebra_def)
-
-lemma (in product_sigma_algebra) sigma_algebra_product_algebra: "sigma_algebra (Pi\<^isub>M I M)"
-  using product_algebra_generator_into_space unfolding product_algebra_def
-  by (rule sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) simp_all
-
-sublocale finite_product_sigma_algebra \<subseteq> sigma_algebra P
-  using sigma_algebra_product_algebra .
+lemma prod_algebra_sets_into_space:
+  "prod_algebra I M \<subseteq> Pow (\<Pi>\<^isub>E i\<in>I. space (M i))"
+  using assms by (auto simp: prod_emb_def prod_algebra_def)
 
-lemma product_algebraE:
-  assumes "A \<in> sets (product_algebra_generator I M)"
-  obtains E where "A = Pi\<^isub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
-  using assms unfolding product_algebra_generator_def by auto
-
-lemma product_algebra_generatorI[intro]:
-  assumes "E \<in> (\<Pi> i\<in>I. sets (M i))"
-  shows "Pi\<^isub>E I E \<in> sets (product_algebra_generator I M)"
-  using assms unfolding product_algebra_generator_def by auto
-
-lemma space_product_algebra_generator[simp]:
-  "space (product_algebra_generator I M) = Pi\<^isub>E I (\<lambda>i. space (M i))"
-  unfolding product_algebra_generator_def by simp
+lemma prod_algebra_eq_finite:
+  assumes I: "finite I"
+  shows "prod_algebra I M = {(\<Pi>\<^isub>E i\<in>I. X i) |X. X \<in> (\<Pi> j\<in>I. sets (M j))}" (is "?L = ?R")
+proof (intro iffI set_eqI)
+  fix A assume "A \<in> ?L"
+  then obtain J E where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
+    and A: "A = prod_emb I M J (PIE j:J. E j)"
+    by (auto simp: prod_algebra_def)
+  let ?A = "\<Pi>\<^isub>E i\<in>I. if i \<in> J then E i else space (M i)"
+  have A: "A = ?A"
+    unfolding A using J by (intro prod_emb_PiE sets_into_space) auto
+  show "A \<in> ?R" unfolding A using J top
+    by (intro CollectI exI[of _ "\<lambda>i. if i \<in> J then E i else space (M i)"]) simp
+next
+  fix A assume "A \<in> ?R"
+  then obtain X where "A = (\<Pi>\<^isub>E i\<in>I. X i)" and X: "X \<in> (\<Pi> j\<in>I. sets (M j))" by auto
+  then have A: "A = prod_emb I M I (\<Pi>\<^isub>E i\<in>I. X i)"
+    using sets_into_space by (force simp: prod_emb_def Pi_iff)
+  from X I show "A \<in> ?L" unfolding A
+    by (auto simp: prod_algebra_def)
+qed
 
-lemma space_product_algebra[simp]:
-  "space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E i\<in>I. space (M i))"
-  unfolding product_algebra_def product_algebra_generator_def by simp
-
-lemma sets_product_algebra:
-  "sets (Pi\<^isub>M I M) = sets (sigma (product_algebra_generator I M))"
-  unfolding product_algebra_def sigma_def by simp
+lemma prod_algebraI:
+  "finite J \<Longrightarrow> (J \<noteq> {} \<or> I = {}) \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i))
+    \<Longrightarrow> prod_emb I M J (PIE j:J. E j) \<in> prod_algebra I M"
+  by (auto simp: prod_algebra_def Pi_iff)
 
-lemma product_algebra_generator_sets_into_space:
-  assumes "\<And>i. i\<in>I \<Longrightarrow> sets (M i) \<subseteq> Pow (space (M i))"
-  shows "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))"
-  using assms by (auto simp: product_algebra_generator_def) blast
-
-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: sets_product_algebra)
-
-lemma Int_stable_product_algebra_generator:
-  "(\<And>i. i \<in> I \<Longrightarrow> Int_stable (M i)) \<Longrightarrow> Int_stable (product_algebra_generator I M)"
-  by (auto simp add: product_algebra_generator_def Int_stable_def PiE_Int Pi_iff)
+lemma prod_algebraE:
+  assumes A: "A \<in> prod_algebra I M"
+  obtains J E where "A = prod_emb I M J (PIE j:J. E j)"
+    "finite J" "J \<noteq> {} \<or> I = {}" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> E i \<in> sets (M i)" 
+  using A by (auto simp: prod_algebra_def)
 
-section "Generating set generates also product algebra"
+lemma prod_algebraE_all:
+  assumes A: "A \<in> prod_algebra I M"
+  obtains E where "A = Pi\<^isub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))"
+proof -
+  from A obtain E J where A: "A = prod_emb I M J (Pi\<^isub>E J E)"
+    and J: "J \<subseteq> I" and E: "E \<in> (\<Pi> i\<in>J. sets (M i))"
+    by (auto simp: prod_algebra_def)
+  from E have "\<And>i. i \<in> J \<Longrightarrow> E i \<subseteq> space (M i)"
+    using sets_into_space by auto
+  then have "A = (\<Pi>\<^isub>E i\<in>I. if i\<in>J then E i else space (M i))"
+    using A J by (auto simp: prod_emb_PiE)
+  moreover then have "(\<lambda>i. if i\<in>J then E i else space (M i)) \<in> (\<Pi> i\<in>I. sets (M i))"
+    using top E by auto
+  ultimately show ?thesis using that by auto
+qed
 
-lemma sigma_product_algebra_sigma_eq:
-  assumes "finite I"
-  assumes mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)"
-  assumes union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = 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 "sets (\<Pi>\<^isub>M i\<in>I. sigma (E i)) = sets (\<Pi>\<^isub>M i\<in>I. E i)"
-    (is "sets ?S = sets ?E")
-proof cases
-  assume "I = {}" then show ?thesis
-    by (simp add: product_algebra_def product_algebra_generator_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 product_algebra_generator_def using E
-    by (intro sigma_algebra.sigma_algebra_cong[OF 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 mono union unfolding incseq_Suc_iff space_product_algebra
-      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 space_product_algebra
-      apply simp
-      apply (subst Pi_UN[OF `finite I`])
-      using mono[THEN incseqD] 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"
-      using sets_into `A \<in> sets (E i)`
-      unfolding sets_product_algebra sets_sigma
-      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: sets_product_algebra 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_generator I (\<lambda>i. sigma (E i)))) \<subseteq> sets ?E"
-    by (intro G.sigma_sets_subset) (auto simp add: product_algebra_generator_def)
-  then have subset: "sets ?S \<subseteq> sets ?E"
-    by (simp add: sets_sigma sets_product_algebra)
-  show "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 sets_product_algebra
-    proof induct
-      case (Basic A) then show ?case
-        by (auto simp: sets_sigma product_algebra_generator_def intro: sigma_sets.Basic)
-    qed (auto intro: sigma_sets.intros simp: product_algebra_generator_def)
-  next
-    fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto
-  qed
+lemma Int_stable_prod_algebra: "Int_stable (prod_algebra I M)"
+proof (unfold Int_stable_def, safe)
+  fix A assume "A \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess J E . note A = this
+  fix B assume "B \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess K F . note B = this
+  have "A \<inter> B = prod_emb I M (J \<union> K) (\<Pi>\<^isub>E i\<in>J \<union> K. (if i \<in> J then E i else space (M i)) \<inter> 
+      (if i \<in> K then F i else space (M i)))"
+    unfolding A B using A(2,3,4) A(5)[THEN sets_into_space] B(2,3,4) B(5)[THEN sets_into_space]
+    apply (subst (1 2 3) prod_emb_PiE)
+    apply (simp_all add: subset_eq PiE_Int)
+    apply blast
+    apply (intro PiE_cong)
+    apply auto
+    done
+  also have "\<dots> \<in> prod_algebra I M"
+    using A B by (auto intro!: prod_algebraI)
+  finally show "A \<inter> B \<in> prod_algebra I M" .
+qed
+
+lemma prod_algebra_mono:
+  assumes space: "\<And>i. i \<in> I \<Longrightarrow> space (E i) = space (F i)"
+  assumes sets: "\<And>i. i \<in> I \<Longrightarrow> sets (E i) \<subseteq> sets (F i)"
+  shows "prod_algebra I E \<subseteq> prod_algebra I F"
+proof
+  fix A assume "A \<in> prod_algebra I E"
+  then obtain J G where J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I"
+    and A: "A = prod_emb I E J (\<Pi>\<^isub>E i\<in>J. G i)"
+    and G: "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (E i)"
+    by (auto simp: prod_algebra_def)
+  moreover
+  from space have "(\<Pi>\<^isub>E i\<in>I. space (E i)) = (\<Pi>\<^isub>E i\<in>I. space (F i))"
+    by (rule PiE_cong)
+  with A have "A = prod_emb I F J (\<Pi>\<^isub>E i\<in>J. G i)"
+    by (simp add: prod_emb_def)
+  moreover
+  from sets G J have "\<And>i. i \<in> J \<Longrightarrow> G i \<in> sets (F i)"
+    by auto
+  ultimately show "A \<in> prod_algebra I F"
+    apply (simp add: prod_algebra_def image_iff)
+    apply (intro exI[of _ J] exI[of _ G] conjI)
+    apply auto
+    done
 qed
 
-lemma product_algebraI[intro]:
-    "E \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> Pi\<^isub>E I E \<in> sets (Pi\<^isub>M I M)"
-  using assms unfolding product_algebra_def by (auto intro: product_algebra_generatorI)
+lemma space_PiM: "space (\<Pi>\<^isub>M i\<in>I. M i) = (\<Pi>\<^isub>E i\<in>I. space (M i))"
+  using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp
+
+lemma sets_PiM: "sets (\<Pi>\<^isub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) (prod_algebra I M)"
+  using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp
 
-lemma (in product_sigma_algebra) measurable_component_update:
-  assumes "x \<in> space (Pi\<^isub>M I M)" and "i \<notin> I"
-  shows "(\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \<in> _")
-  unfolding product_algebra_def apply simp
-proof (intro measurable_sigma)
-  let ?G = "product_algebra_generator (insert i I) M"
-  show "sets ?G \<subseteq> Pow (space ?G)" using product_algebra_generator_into_space .
-  show "?f \<in> space (M i) \<rightarrow> space ?G"
-    using M.sets_into_space assms by auto
-  fix A assume "A \<in> sets ?G"
-  from product_algebraE[OF this] guess E . note E = this
-  then have "?f -` A \<inter> space (M i) = E i \<or> ?f -` A \<inter> space (M i) = {}"
-    using M.sets_into_space assms by auto
-  then show "?f -` A \<inter> space (M i) \<in> sets (M i)"
-    using E by (auto intro!: product_algebraI)
+lemma sets_PiM_single: "sets (PiM I M) =
+    sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) {{f\<in>\<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} | i A. i \<in> I \<and> A \<in> sets (M i)}"
+    (is "_ = sigma_sets ?\<Omega> ?R")
+  unfolding sets_PiM
+proof (rule sigma_sets_eqI)
+  interpret R: sigma_algebra ?\<Omega> "sigma_sets ?\<Omega> ?R" by (rule sigma_algebra_sigma_sets) auto
+  fix A assume "A \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess J X . note X = this
+  show "A \<in> sigma_sets ?\<Omega> ?R"
+  proof cases
+    assume "I = {}"
+    with X have "A = {\<lambda>x. undefined}" by (auto simp: prod_emb_def)
+    with `I = {}` show ?thesis by (auto intro!: sigma_sets_top)
+  next
+    assume "I \<noteq> {}"
+    with X have "A = (\<Inter>j\<in>J. {f\<in>(\<Pi>\<^isub>E i\<in>I. space (M i)). f j \<in> X j})"
+      using sets_into_space[OF X(5)]
+      by (auto simp: prod_emb_PiE[OF _ sets_into_space] Pi_iff split: split_if_asm) blast
+    also have "\<dots> \<in> sigma_sets ?\<Omega> ?R"
+      using X `I \<noteq> {}` by (intro R.finite_INT sigma_sets.Basic) auto
+    finally show "A \<in> sigma_sets ?\<Omega> ?R" .
+  qed
+next
+  fix A assume "A \<in> ?R"
+  then obtain i B where A: "A = {f\<in>\<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> B}" "i \<in> I" "B \<in> sets (M i)" 
+    by auto
+  then have "A = prod_emb I M {i} (\<Pi>\<^isub>E i\<in>{i}. B)"
+    using sets_into_space[OF A(3)]
+    apply (subst prod_emb_PiE)
+    apply (auto simp: Pi_iff split: split_if_asm)
+    apply blast
+    done
+  also have "\<dots> \<in> sigma_sets ?\<Omega> (prod_algebra I M)"
+    using A by (intro sigma_sets.Basic prod_algebraI) auto
+  finally show "A \<in> sigma_sets ?\<Omega> (prod_algebra I M)" .
+qed
+
+lemma sets_PiM_I:
+  assumes "finite J" "J \<subseteq> I" "\<forall>i\<in>J. E i \<in> sets (M i)"
+  shows "prod_emb I M J (PIE j:J. E j) \<in> sets (PIM i:I. M i)"
+proof cases
+  assume "J = {}"
+  then have "prod_emb I M J (PIE j:J. E j) = (PIE j:I. space (M j))"
+    by (auto simp: prod_emb_def)
+  then show ?thesis
+    by (auto simp add: sets_PiM intro!: sigma_sets_top)
+next
+  assume "J \<noteq> {}" with assms show ?thesis
+    by (auto simp add: sets_PiM prod_algebra_def intro!: sigma_sets.Basic)
 qed
 
-lemma (in product_sigma_algebra) measurable_add_dim:
-  assumes "i \<notin> I"
-  shows "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)"
-proof -
-  let ?f = "(\<lambda>(f, y). f(i := y))" and ?G = "product_algebra_generator (insert i I) M"
-  interpret Ii: pair_sigma_algebra "Pi\<^isub>M I M" "M i"
-    unfolding pair_sigma_algebra_def
-    by (intro sigma_algebra_product_algebra sigma_algebras conjI)
-  have "?f \<in> measurable Ii.P (sigma ?G)"
-  proof (rule Ii.measurable_sigma)
-    show "sets ?G \<subseteq> Pow (space ?G)"
-      using product_algebra_generator_into_space .
-    show "?f \<in> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<rightarrow> space ?G"
-      by (auto simp: space_pair_measure)
-  next
-    fix A assume "A \<in> sets ?G"
-    then obtain F where "A = Pi\<^isub>E (insert i I) F"
-      and F: "\<And>j. j \<in> I \<Longrightarrow> F j \<in> sets (M j)" "F i \<in> sets (M i)"
-      by (auto elim!: product_algebraE)
-    then have "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = Pi\<^isub>E I F \<times> (F i)"
-      using sets_into_space `i \<notin> I`
-      by (auto simp add: space_pair_measure) blast+
-    then show "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M M i)"
-      using F by (auto intro!: pair_measureI)
-  qed
-  then show ?thesis
-    by (simp add: product_algebra_def)
+lemma measurable_PiM:
+  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
+  assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
+    f -` prod_emb I M J (Pi\<^isub>E J X) \<inter> space N \<in> sets N" 
+  shows "f \<in> measurable N (PiM I M)"
+  using sets_PiM prod_algebra_sets_into_space space
+proof (rule measurable_sigma_sets)
+  fix A assume "A \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess J X .
+  with sets[of J X] show "f -` A \<inter> space N \<in> sets N" by auto
+qed
+
+lemma measurable_PiM_Collect:
+  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
+  assumes sets: "\<And>X J. J \<noteq> {} \<or> I = {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow>
+    {\<omega>\<in>space N. \<forall>i\<in>J. f \<omega> i \<in> X i} \<in> sets N" 
+  shows "f \<in> measurable N (PiM I M)"
+  using sets_PiM prod_algebra_sets_into_space space
+proof (rule measurable_sigma_sets)
+  fix A assume "A \<in> prod_algebra I M"
+  from prod_algebraE[OF this] guess J X . note X = this
+  have "f -` A \<inter> space N = {\<omega> \<in> space N. \<forall>i\<in>J. f \<omega> i \<in> X i}"
+    using sets_into_space[OF X(5)] X(2-) space unfolding X(1)
+    by (subst prod_emb_PiE) (auto simp: Pi_iff split: split_if_asm)
+  also have "\<dots> \<in> sets N" using X(3,2,4,5) by (rule sets)
+  finally show "f -` A \<inter> space N \<in> sets N" .
 qed
 
-lemma (in product_sigma_algebra) measurable_merge:
-  assumes [simp]: "I \<inter> J = {}"
-  shows "(\<lambda>(x, y). merge I x J y) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
-proof -
-  let ?I = "Pi\<^isub>M I M" and ?J = "Pi\<^isub>M J M"
-  interpret P: sigma_algebra "?I \<Otimes>\<^isub>M ?J"
-    by (intro sigma_algebra_pair_measure product_algebra_into_space)
-  let ?f = "\<lambda>(x, y). merge I x J y"
-  let ?G = "product_algebra_generator (I \<union> J) M"
-  have "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (sigma ?G)"
-  proof (rule P.measurable_sigma)
-    fix A assume "A \<in> sets ?G"
-    from product_algebraE[OF this]
-    obtain E where E: "A = Pi\<^isub>E (I \<union> J) E" "E \<in> (\<Pi> i\<in>I \<union> J. sets (M i))" .
-    then have *: "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) = Pi\<^isub>E I E \<times> Pi\<^isub>E J E"
-      using sets_into_space `I \<inter> J = {}`
-      by (auto simp add: space_pair_measure) fast+
-    then show "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) \<in> sets (?I \<Otimes>\<^isub>M ?J)"
-      using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI
-        simp: product_algebra_def)
-  qed (insert product_algebra_generator_into_space, auto simp: space_pair_measure)
-  then show "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (Pi\<^isub>M (I \<union> J) M)"
-    unfolding product_algebra_def[of "I \<union> J"] by simp
-qed
+lemma measurable_PiM_single:
+  assumes space: "f \<in> space N \<rightarrow> (\<Pi>\<^isub>E i\<in>I. space (M i))"
+  assumes sets: "\<And>A i. i \<in> I \<Longrightarrow> A \<in> sets (M i) \<Longrightarrow> {\<omega> \<in> space N. f \<omega> i \<in> A} \<in> sets N" 
+  shows "f \<in> measurable N (PiM I M)"
+  using sets_PiM_single
+proof (rule measurable_sigma_sets)
+  fix A assume "A \<in> {{f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
+  then obtain B i where "A = {f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> B}" and B: "i \<in> I" "B \<in> sets (M i)"
+    by auto
+  with space have "f -` A \<inter> space N = {\<omega> \<in> space N. f \<omega> i \<in> B}" by auto
+  also have "\<dots> \<in> sets N" using B by (rule sets)
+  finally show "f -` A \<inter> space N \<in> sets N" .
+qed (auto simp: space)
 
-lemma (in product_sigma_algebra) measurable_component_singleton:
+lemma sets_PiM_I_finite[simp, intro]:
+  assumes "finite I" and sets: "(\<And>i. i \<in> I \<Longrightarrow> E i \<in> sets (M i))"
+  shows "(PIE j:I. E j) \<in> sets (PIM i:I. M i)"
+  using sets_PiM_I[of I I E M] sets_into_space[OF sets] `finite I` sets by auto
+
+lemma measurable_component_update:
+  assumes "x \<in> space (Pi\<^isub>M I M)" and "i \<notin> I"
+  shows "(\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \<in> _")
+proof (intro measurable_PiM_single)
+  fix j A assume "j \<in> insert i I" "A \<in> sets (M j)"
+  moreover have "{\<omega> \<in> space (M i). (x(i := \<omega>)) j \<in> A} =
+    (if i = j then space (M i) \<inter> A else if x j \<in> A then space (M i) else {})"
+    by auto
+  ultimately show "{\<omega> \<in> space (M i). (x(i := \<omega>)) j \<in> A} \<in> sets (M i)"
+    by auto
+qed (insert sets_into_space assms, auto simp: space_PiM)
+
+lemma measurable_component_singleton:
   assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>M I M) (M i)"
 proof (unfold measurable_def, intro CollectI conjI ballI)
   fix A assume "A \<in> sets (M i)"
-  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))"
-    using M.sets_into_space `i \<in> I` by (fastforce dest: Pi_mem split: split_if_asm)
+  then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) = prod_emb I M {i} (\<Pi>\<^isub>E j\<in>{i}. A)"
+    using sets_into_space `i \<in> I`
+    by (fastforce dest: Pi_mem simp: prod_emb_def space_PiM split: split_if_asm)
   then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) \<in> sets (Pi\<^isub>M I M)"
-    using `A \<in> sets (M i)` by (auto intro!: product_algebraI)
-qed (insert `i \<in> I`, auto)
+    using `A \<in> sets (M i)` `i \<in> I` by (auto intro!: sets_PiM_I)
+qed (insert `i \<in> I`, auto simp: space_PiM)
+
+lemma measurable_add_dim:
+  assumes "i \<notin> I"
+  shows "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)"
+    (is "?f \<in> measurable ?P ?I")
+proof (rule measurable_PiM_single)
+  fix j A assume j: "j \<in> insert i I" and A: "A \<in> sets (M j)"
+  have "{\<omega> \<in> space ?P. (\<lambda>(f, x). fun_upd f i x) \<omega> j \<in> A} =
+    (if j = i then space (Pi\<^isub>M I M) \<times> A else ((\<lambda>x. x j) \<circ> fst) -` A \<inter> space ?P)"
+    using sets_into_space[OF A] by (auto simp add: space_pair_measure space_PiM)
+  also have "\<dots> \<in> sets ?P"
+    using A j
+    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
+  finally show "{\<omega> \<in> space ?P. prod_case (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" .
+qed (auto simp: space_pair_measure space_PiM)
 
-lemma (in sigma_algebra) measurable_restrict:
-  assumes I: "finite I"
-  assumes "\<And>i. i \<in> I \<Longrightarrow> sets (N i) \<subseteq> Pow (space (N i))"
-  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable M (N i)"
-  shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable M (Pi\<^isub>M I N)"
-  unfolding product_algebra_def
-proof (simp, rule measurable_sigma)
-  show "sets (product_algebra_generator I N) \<subseteq> Pow (space (product_algebra_generator I N))"
-    by (rule product_algebra_generator_sets_into_space) fact
-  show "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> space M \<rightarrow> space (product_algebra_generator I N)"
-    using X by (auto simp: measurable_def)
-  fix E assume "E \<in> sets (product_algebra_generator I N)"
-  then obtain F where "E = Pi\<^isub>E I F" and F: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets (N i)"
-    by (auto simp: product_algebra_generator_def)
-  then have "(\<lambda>x. \<lambda>i\<in>I. X i x) -` E \<inter> space M = (\<Inter>i\<in>I. X i -` F i \<inter> space M) \<inter> space M"
-    by (auto simp: Pi_iff)
-  also have "\<dots> \<in> sets M"
-  proof cases
-    assume "I = {}" then show ?thesis by simp
-  next
-    assume "I \<noteq> {}" with X F I show ?thesis
-      by (intro finite_INT measurable_sets Int) auto
-  qed
-  finally show "(\<lambda>x. \<lambda>i\<in>I. X i x) -` E \<inter> space M \<in> sets M" .
-qed
+lemma measurable_merge:
+  assumes "I \<inter> J = {}"
+  shows "(\<lambda>(x, y). merge I x J y) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
+    (is "?f \<in> measurable ?P ?U")
+proof (rule measurable_PiM_single)
+  fix i A assume A: "A \<in> sets (M i)" "i \<in> I \<union> J"
+  then have "{\<omega> \<in> space ?P. prod_case (\<lambda>x. merge I x J) \<omega> i \<in> A} =
+    (if i \<in> I then ((\<lambda>x. x i) \<circ> fst) -` A \<inter> space ?P else ((\<lambda>x. x i) \<circ> snd) -` A \<inter> space ?P)"
+    using `I \<inter> J = {}` by auto
+  also have "\<dots> \<in> sets ?P"
+    using A
+    by (auto intro!: measurable_sets[OF measurable_comp, OF _ measurable_component_singleton])
+  finally show "{\<omega> \<in> space ?P. prod_case (\<lambda>x. merge I x J) \<omega> i \<in> A} \<in> sets ?P" .
+qed (insert assms, auto simp: space_pair_measure space_PiM)
 
-locale product_sigma_finite = product_sigma_algebra M
-  for M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" +
+lemma measurable_restrict:
+  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable N (M i)"
+  shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable N (Pi\<^isub>M I M)"
+proof (rule measurable_PiM_single)
+  fix A i assume A: "i \<in> I" "A \<in> sets (M i)"
+  then have "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} = X i -` A \<inter> space N"
+    by auto
+  then show "{\<omega> \<in> space N. (\<lambda>i\<in>I. X i \<omega>) i \<in> A} \<in> sets N"
+    using A X by (auto intro!: measurable_sets)
+qed (insert X, auto dest: measurable_space)
+
+locale product_sigma_finite =
+  fixes M :: "'i \<Rightarrow> 'a measure"
   assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)"
 
 sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" for i
   by (rule sigma_finite_measures)
 
-locale finite_product_sigma_finite = finite_product_sigma_algebra M I + product_sigma_finite M
-  for M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" and I :: "'i set"
-
-lemma (in finite_product_sigma_finite) product_algebra_generator_measure:
-  assumes "Pi\<^isub>E I F \<in> sets G"
-  shows "measure G (Pi\<^isub>E I F) = (\<Prod>i\<in>I. M.\<mu> i (F i))"
-proof cases
-  assume ne: "\<forall>i\<in>I. F i \<noteq> {}"
-  have "\<forall>i\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') i = F i"
-    by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
-       (insert ne, auto simp: Pi_eq_iff)
-  then show ?thesis
-    unfolding product_algebra_generator_def by simp
-next
-  assume empty: "\<not> (\<forall>j\<in>I. F j \<noteq> {})"
-  then have "(\<Prod>j\<in>I. M.\<mu> j (F j)) = 0"
-    by (auto simp: setprod_ereal_0 intro!: bexI)
-  moreover
-  have "\<exists>j\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}"
-    by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"])
-       (insert empty, auto simp: Pi_eq_empty_iff[symmetric])
-  then have "(\<Prod>j\<in>I. M.\<mu> j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0"
-    by (auto simp: setprod_ereal_0 intro!: bexI)
-  ultimately show ?thesis
-    unfolding product_algebra_generator_def by simp
-qed
+locale finite_product_sigma_finite = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
+  fixes I :: "'i set"
+  assumes finite_index: "finite I"
 
 lemma (in finite_product_sigma_finite) sigma_finite_pairs:
   "\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set.
     (\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and>
-    (\<forall>k. \<forall>i\<in>I. \<mu> i (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<and>
-    (\<Union>k. \<Pi>\<^isub>E i\<in>I. F i k) = space G"
+    (\<forall>k. \<forall>i\<in>I. emeasure (M i) (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<and>
+    (\<Union>k. \<Pi>\<^isub>E i\<in>I. F i k) = space (PiM I M)"
 proof -
-  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. \<mu> i (F k) \<noteq> \<infinity>)"
-    using M.sigma_finite_up by simp
+  have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. emeasure (M i) (F k) \<noteq> \<infinity>)"
+    using M.sigma_finite_incseq by metis
   from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
-  then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. \<mu> i (F i k) \<noteq> \<infinity>"
+  then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. emeasure (M i) (F i k) \<noteq> \<infinity>"
     by auto
   let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k"
-  note space_product_algebra[simp]
+  note space_PiM[simp]
   show ?thesis
   proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI)
     fix i show "range (F i) \<subseteq> sets (M i)" by fact
   next
-    fix i k show "\<mu> i (F i k) \<noteq> \<infinity>" by fact
+    fix i k show "emeasure (M i) (F i k) \<noteq> \<infinity>" 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 (force simp: image_subset_iff)
+    fix A assume "A \<in> (\<Union>i. ?F i)" then show "A \<in> space (PiM I M)"
+      using `\<And>i. range (F i) \<subseteq> sets (M i)` sets_into_space
+      by auto blast
   next
-    fix f assume "f \<in> space G"
+    fix f assume "f \<in> space (PiM I M)"
     with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F
     show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def)
   next
@@ -574,140 +590,144 @@
   qed
 qed
 
-lemma sets_pair_cancel_measure[simp]:
-  "sets (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) = sets (M1 \<Otimes>\<^isub>M M2)"
-  "sets (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) = sets (M1 \<Otimes>\<^isub>M M2)"
-  unfolding pair_measure_def pair_measure_generator_def sets_sigma
-  by simp_all
-
-lemma measurable_pair_cancel_measure[simp]:
-  "measurable (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) M = measurable (M1 \<Otimes>\<^isub>M M2) M"
-  "measurable (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) M = measurable (M1 \<Otimes>\<^isub>M M2) M"
-  "measurable M (M1\<lparr>measure := m3\<rparr> \<Otimes>\<^isub>M M2) = measurable M (M1 \<Otimes>\<^isub>M M2)"
-  "measurable M (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m4\<rparr>) = measurable M (M1 \<Otimes>\<^isub>M M2)"
-  unfolding measurable_def by (auto simp add: space_pair_measure)
-
-lemma (in product_sigma_finite) product_measure_exists:
+lemma (in product_sigma_finite)
   assumes "finite I"
-  shows "\<exists>\<nu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<nu> \<rparr>) \<and>
-    (\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i)))"
+  shows sigma_finite: "sigma_finite_measure (PiM I M)"
+  and emeasure_PiM:
+    "\<And>A. (\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (PiM I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
 using `finite I` proof induct
   case empty
-  interpret finite_product_sigma_finite M "{}" by default simp
-  let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> ereal"
-  show ?case
-  proof (intro exI conjI ballI)
-    have "sigma_algebra (sigma G \<lparr>measure := ?\<nu>\<rparr>)"
-      by (rule sigma_algebra_cong) (simp_all add: product_algebra_def)
-    then have "measure_space (sigma G\<lparr>measure := ?\<nu>\<rparr>)"
-      by (rule finite_additivity_sufficient)
-         (simp_all add: positive_def additive_def sets_sigma
-                        product_algebra_generator_def image_constant)
-    then show "sigma_finite_measure (sigma G\<lparr>measure := ?\<nu>\<rparr>)"
-      by (auto intro!: exI[of _ "\<lambda>i. {\<lambda>_. undefined}"]
-               simp: sigma_finite_measure_def sigma_finite_measure_axioms_def
-                     product_algebra_generator_def)
-  qed auto
+  let ?\<mu> = "\<lambda>A. if A = {} then 0 else (1::ereal)"
+  have "prod_algebra {} M = {{\<lambda>_. undefined}}"
+    by (auto simp: prod_algebra_def prod_emb_def intro!: image_eqI)
+  then have sets_empty: "sets (PiM {} M) = {{}, {\<lambda>_. undefined}}"
+    by (simp add: sets_PiM)
+  have "emeasure (Pi\<^isub>M {} M) (prod_emb {} M {} (\<Pi>\<^isub>E i\<in>{}. {})) = 1"
+  proof (subst emeasure_extend_measure_Pair[OF PiM_def])
+    have "finite (space (PiM {} M))"
+      by (simp add: space_PiM)
+    moreover show "positive (PiM {} M) ?\<mu>"
+      by (auto simp: positive_def)
+    ultimately show "countably_additive (PiM {} M) ?\<mu>"
+      by (rule countably_additiveI_finite) (auto simp: additive_def space_PiM sets_empty)
+  qed (auto simp: prod_emb_def)
+  also have *: "(prod_emb {} M {} (\<Pi>\<^isub>E i\<in>{}. {})) = {\<lambda>_. undefined}"
+    by (auto simp: prod_emb_def)
+  finally have emeasure_eq: "emeasure (Pi\<^isub>M {} M) {\<lambda>_. undefined} = 1" .
+
+  interpret finite_measure "PiM {} M"
+    by default (simp add: space_PiM emeasure_eq)
+  case 1 show ?case ..
+
+  case 2 show ?case
+    using emeasure_eq * by simp
 next
   case (insert i I)
   interpret finite_product_sigma_finite M I by default fact
   have "finite (insert i I)" using `finite I` by auto
   interpret I': finite_product_sigma_finite M "insert i I" by default fact
-  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. M.\<mu> i (A i))" and
-    "sigma_finite_measure (sigma G\<lparr> measure := \<nu> \<rparr>)" by auto
-  then interpret I: sigma_finite_measure "P\<lparr> measure := \<nu>\<rparr>" unfolding product_algebra_def by simp
-  interpret P: pair_sigma_finite "P\<lparr> measure := \<nu>\<rparr>" "M i" ..
+  interpret I: sigma_finite_measure "PiM I M" by fact
+  interpret P: pair_sigma_finite "PiM I M" "M i" ..
   let ?h = "(\<lambda>(f, y). f(i := y))"
-  let ?\<nu> = "\<lambda>A. P.\<mu> (?h -` A \<inter> space P.P)"
-  have I': "sigma_algebra (I'.P\<lparr> measure := ?\<nu> \<rparr>)"
-    by (rule I'.sigma_algebra_cong) simp_all
-  interpret I'': measure_space "I'.P\<lparr> measure := ?\<nu> \<rparr>"
-    using measurable_add_dim[OF `i \<notin> I`]
-    by (intro P.measure_space_vimage[OF I']) (auto simp add: measure_preserving_def)
-  show ?case
-  proof (intro exI[of _ ?\<nu>] conjI ballI)
-    let ?m = "\<lambda>A. measure (Pi\<^isub>M I M\<lparr>measure := \<nu>\<rparr> \<Otimes>\<^isub>M M i) (?h -` A \<inter> space P.P)"
-    { fix A assume A: "A \<in> (\<Pi> i\<in>insert i I. sets (M i))"
-      then have *: "?h -` Pi\<^isub>E (insert i I) A \<inter> space P.P = Pi\<^isub>E I A \<times> A i"
-        using `i \<notin> I` M.sets_into_space by (auto simp: space_pair_measure space_product_algebra) blast
-      show "?m (Pi\<^isub>E (insert i I) A) = (\<Prod>i\<in>insert i I. M.\<mu> i (A i))"
-        unfolding * using A
-        apply (subst P.pair_measure_times)
-        using A apply fastforce
-        using A apply fastforce
-        using `i \<notin> I` `finite I` prod[of A] A by (auto simp: ac_simps) }
-    note product = this
-    have *: "sigma I'.G\<lparr> measure := ?\<nu> \<rparr> = I'.P\<lparr> measure := ?\<nu> \<rparr>"
-      by (simp add: product_algebra_def)
-    show "sigma_finite_measure (sigma I'.G\<lparr> measure := ?\<nu> \<rparr>)"
-    proof (unfold *, default, simp)
-      from I'.sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
-      then have F: "\<And>j. j \<in> insert i I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
-        "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k)"
-        "(\<Union>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) = space I'.G"
-        "\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> \<mu> j (F j k) \<noteq> \<infinity>"
-        by blast+
-      let ?F = "\<lambda>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) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i)) \<and> (\<forall>i. ?m (F i) \<noteq> \<infinity>)"
-      proof (intro exI[of _ ?F] conjI allI)
-        show "range ?F \<subseteq> sets I'.P" using F(1) by auto
-      next
-        from F(3) show "(\<Union>i. ?F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i))" by simp
-      next
-        fix j
-        have "\<And>k. k \<in> insert i I \<Longrightarrow> 0 \<le> measure (M k) (F k j)"
-          using F(1) by auto
-        with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\<nu> (?F j) \<noteq> \<infinity>"
-          by (subst product) auto
-      qed
-    qed
+
+  let ?P = "distr (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M) ?h"
+  let ?\<mu> = "emeasure ?P"
+  let ?I = "{j \<in> insert i I. emeasure (M j) (space (M j)) \<noteq> 1}"
+  let ?f = "\<lambda>J E j. if j \<in> J then emeasure (M j) (E j) else emeasure (M j) (space (M j))"
+
+  { case 2
+    have "emeasure (Pi\<^isub>M (insert i I) M) (prod_emb (insert i I) M (insert i I) (Pi\<^isub>E (insert i I) A)) =
+      (\<Prod>i\<in>insert i I. emeasure (M i) (A i))"
+    proof (subst emeasure_extend_measure_Pair[OF PiM_def])
+      fix J E assume "(J \<noteq> {} \<or> insert i I = {}) \<and> finite J \<and> J \<subseteq> insert i I \<and> E \<in> (\<Pi> j\<in>J. sets (M j))"
+      then have J: "J \<noteq> {}" "finite J" "J \<subseteq> insert i I" and E: "\<forall>j\<in>J. E j \<in> sets (M j)" by auto
+      let ?p = "prod_emb (insert i I) M J (Pi\<^isub>E J E)"
+      let ?p' = "prod_emb I M (J - {i}) (\<Pi>\<^isub>E j\<in>J-{i}. E j)"
+      have "?\<mu> ?p =
+        emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i))"
+        by (intro emeasure_distr measurable_add_dim sets_PiM_I) fact+
+      also have "?h -` ?p \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = ?p' \<times> (if i \<in> J then E i else space (M i))"
+        using J E[rule_format, THEN sets_into_space]
+        by (force simp: space_pair_measure space_PiM Pi_iff prod_emb_iff split: split_if_asm)
+      also have "emeasure (Pi\<^isub>M I M \<Otimes>\<^isub>M (M i)) (?p' \<times> (if i \<in> J then E i else space (M i))) =
+        emeasure (Pi\<^isub>M I M) ?p' * emeasure (M i) (if i \<in> J then (E i) else space (M i))"
+        using J E by (intro P.emeasure_pair_measure_Times sets_PiM_I) auto
+      also have "?p' = (\<Pi>\<^isub>E j\<in>I. if j \<in> J-{i} then E j else space (M j))"
+        using J E[rule_format, THEN sets_into_space]
+        by (auto simp: prod_emb_iff Pi_iff split: split_if_asm) blast+
+      also have "emeasure (Pi\<^isub>M I M) (\<Pi>\<^isub>E j\<in>I. if j \<in> J-{i} then E j else space (M j)) =
+        (\<Prod> j\<in>I. if j \<in> J-{i} then emeasure (M j) (E j) else emeasure (M j) (space (M j)))"
+        using E by (subst insert) (auto intro!: setprod_cong)
+      also have "(\<Prod>j\<in>I. if j \<in> J - {i} then emeasure (M j) (E j) else emeasure (M j) (space (M j))) *
+         emeasure (M i) (if i \<in> J then E i else space (M i)) = (\<Prod>j\<in>insert i I. ?f J E j)"
+        using insert by (auto simp: mult_commute intro!: arg_cong2[where f="op *"] setprod_cong)
+      also have "\<dots> = (\<Prod>j\<in>J \<union> ?I. ?f J E j)"
+        using insert(1,2) J E by (intro setprod_mono_one_right) auto
+      finally show "?\<mu> ?p = \<dots>" .
+
+      show "prod_emb (insert i I) M J (Pi\<^isub>E J E) \<in> Pow (\<Pi>\<^isub>E i\<in>insert i I. space (M i))"
+        using J E[rule_format, THEN sets_into_space] by (auto simp: prod_emb_iff)
+    next
+      show "positive (sets (Pi\<^isub>M (insert i I) M)) ?\<mu>" "countably_additive (sets (Pi\<^isub>M (insert i I) M)) ?\<mu>"
+        using emeasure_positive[of ?P] emeasure_countably_additive[of ?P] by simp_all
+    next
+      show "(insert i I \<noteq> {} \<or> insert i I = {}) \<and> finite (insert i I) \<and>
+        insert i I \<subseteq> insert i I \<and> A \<in> (\<Pi> j\<in>insert i I. sets (M j))"
+        using insert(1,2) 2 by auto
+    qed (auto intro!: setprod_cong)
+    with 2[THEN sets_into_space] show ?case by (subst (asm) prod_emb_PiE_same_index) auto }
+  note product = this
+
+  from I'.sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" ..
+  then have F: "\<And>j. j \<in> insert i I \<Longrightarrow> range (F j) \<subseteq> sets (M j)"
+    "incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k)"
+    "(\<Union>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) = space (Pi\<^isub>M (insert i I) M)"
+    "\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> emeasure (M j) (F j k) \<noteq> \<infinity>"
+    by blast+
+  let ?F = "\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k"
+
+  case 1 show ?case
+  proof (unfold_locales, intro exI[of _ ?F] conjI allI)
+    show "range ?F \<subseteq> sets (Pi\<^isub>M (insert i I) M)" using F(1) insert(1,2) by auto
+  next
+    from F(3) show "(\<Union>i. ?F i) = space (Pi\<^isub>M (insert i I) M)" by simp
+  next
+    fix j
+    from F `finite I` setprod_PInf[of "insert i I", OF emeasure_nonneg, of M]
+    show "emeasure (\<Pi>\<^isub>M i\<in>insert i I. M i) (?F j) \<noteq> \<infinity>"
+      by (subst product) auto
   qed
 qed
 
-sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure P
-  unfolding product_algebra_def
-  using product_measure_exists[OF finite_index]
-  by (rule someI2_ex) auto
+sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure "Pi\<^isub>M I M"
+  using sigma_finite[OF finite_index] .
 
 lemma (in finite_product_sigma_finite) measure_times:
-  assumes "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)"
-  shows "\<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
-  unfolding product_algebra_def
-  using product_measure_exists[OF finite_index]
-  proof (rule someI2_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. M.\<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. M.\<mu> i ((\<lambda>i\<in>I. A i) i))" using assms * by auto
-    finally show "measure (sigma G\<lparr>measure := \<nu>\<rparr>) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))"
-      by (simp add: sigma_def)
-  qed
+  "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> emeasure (Pi\<^isub>M I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
+  using emeasure_PiM[OF finite_index] by auto
 
 lemma (in product_sigma_finite) product_measure_empty[simp]:
-  "measure (Pi\<^isub>M {} M) {\<lambda>x. undefined} = 1"
+  "emeasure (Pi\<^isub>M {} M) {\<lambda>x. undefined} = 1"
 proof -
   interpret finite_product_sigma_finite M "{}" by default auto
   from measure_times[of "\<lambda>x. {}"] show ?thesis by simp
 qed
 
-lemma (in finite_product_sigma_algebra) P_empty:
-  assumes "I = {}"
-  shows "space P = {\<lambda>k. undefined}" "sets P = { {}, {\<lambda>k. undefined} }"
-  unfolding product_algebra_def product_algebra_generator_def `I = {}`
-  by (simp_all add: sigma_def image_constant)
+lemma
+  shows space_PiM_empty: "space (Pi\<^isub>M {} M) = {\<lambda>k. undefined}"
+    and sets_PiM_empty: "sets (Pi\<^isub>M {} M) = { {}, {\<lambda>k. undefined} }"
+  by (simp_all add: space_PiM sets_PiM_single image_constant sigma_sets_empty_eq)
 
 lemma (in product_sigma_finite) positive_integral_empty:
   assumes pos: "0 \<le> f (\<lambda>k. undefined)"
   shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)"
 proof -
   interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI)
-  have "\<And>A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1"
+  have "\<And>A. emeasure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1"
     using assms by (subst measure_times) auto
   then show ?thesis
-    unfolding positive_integral_def simple_function_def simple_integral_def [abs_def]
-  proof (simp add: P_empty, intro antisym)
+    unfolding positive_integral_def simple_function_def simple_integral_def[abs_def]
+  proof (simp add: space_PiM_empty sets_PiM_empty, intro antisym)
     show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined))"
       by (intro SUP_upper) (auto simp: le_fun_def split: split_max)
     show "(SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)" using pos
@@ -715,71 +735,72 @@
   qed
 qed
 
-lemma (in product_sigma_finite) measure_fold:
+lemma (in product_sigma_finite) distr_merge:
   assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
-  assumes A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
-  shows "measure (Pi\<^isub>M (I \<union> J) M) A =
-    measure (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ((\<lambda>(x,y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M))"
+  shows "distr (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M) (\<lambda>(x,y). merge I x J y) = Pi\<^isub>M (I \<union> J) M"
+   (is "?D = ?P")
 proof -
   interpret I: finite_product_sigma_finite M I by default fact
   interpret J: finite_product_sigma_finite M J by default fact
   have "finite (I \<union> J)" using fin by auto
   interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
-  interpret P: pair_sigma_finite I.P J.P by default
+  interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
   let ?g = "\<lambda>(x,y). merge I x J y"
-  let ?X = "\<lambda>S. ?g -` S \<inter> space P.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)"
        "incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k)"
-       "(\<Union>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) = space IJ.G"
-       "\<And>k. \<forall>i\<in>I\<union>J. \<mu> i (F i k) \<noteq> \<infinity>"
+       "(\<Union>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) = space ?P"
+       "\<And>k. \<forall>i\<in>I\<union>J. emeasure (M i) (F i k) \<noteq> \<infinity>"
     by auto
   let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k"
-  show "IJ.\<mu> A = P.\<mu> (?X A)"
-  proof (rule measure_unique_Int_stable_vimage)
-    show "measure_space IJ.P" "measure_space P.P" by default
-    show "sets (sigma IJ.G) = sets IJ.P" "space IJ.G = space IJ.P" "A \<in> sets (sigma IJ.G)"
-      using A unfolding product_algebra_def by auto
-  next
-    show "Int_stable IJ.G"
-      by (rule Int_stable_product_algebra_generator)
-         (auto simp: Int_stable_def)
-    show "range ?F \<subseteq> sets IJ.G" using F
-      by (simp add: image_subset_iff product_algebra_def
-                    product_algebra_generator_def)
-    show "incseq ?F" "(\<Union>i. ?F i) = space IJ.G " by fact+
+  
+  show ?thesis
+  proof (rule measure_eqI_generator_eq[symmetric])
+    show "Int_stable (prod_algebra (I \<union> J) M)"
+      by (rule Int_stable_prod_algebra)
+    show "prod_algebra (I \<union> J) M \<subseteq> Pow (\<Pi>\<^isub>E i \<in> I \<union> J. space (M i))"
+      by (rule prod_algebra_sets_into_space)
+    show "sets ?P = sigma_sets (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
+      by (rule sets_PiM)
+    then show "sets ?D = sigma_sets (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
+      by simp
+
+    show "range ?F \<subseteq> prod_algebra (I \<union> J) M" using F
+      using fin by (auto simp: prod_algebra_eq_finite)
+    show "incseq ?F" by fact
+    show "(\<Union>i. \<Pi>\<^isub>E ia\<in>I \<union> J. F ia i) = (\<Pi>\<^isub>E i\<in>I \<union> J. space (M i))"
+      using F(3) by (simp add: space_PiM)
   next
     fix k
-    have "\<And>j. j \<in> I \<union> J \<Longrightarrow> 0 \<le> measure (M j) (F j k)"
-      using F(1) by auto
-    with F `finite I` setprod_PInf[of "I \<union> J", OF this]
-    show "IJ.\<mu> (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto
+    from F `finite I` setprod_PInf[of "I \<union> J", OF emeasure_nonneg, of M]
+    show "emeasure ?P (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto
   next
-    fix A assume "A \<in> sets IJ.G"
-    then obtain F where A: "A = Pi\<^isub>E (I \<union> J) F"
-      and F: "\<And>i. i \<in> I \<union> J \<Longrightarrow> F i \<in> sets (M i)"
-      by (auto simp: product_algebra_generator_def)
-    then have X: "?X A = (Pi\<^isub>E I F \<times> Pi\<^isub>E J F)"
-      using sets_into_space by (auto simp: space_pair_measure) blast+
-    then have "P.\<mu> (?X 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))"
+    fix A assume A: "A \<in> prod_algebra (I \<union> J) M"
+    with fin obtain F where A_eq: "A = (Pi\<^isub>E (I \<union> J) F)" and F: "\<forall>i\<in>I \<union> J. F i \<in> sets (M i)"
+      by (auto simp add: prod_algebra_eq_finite)
+    let ?B = "Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M"
+    let ?X = "?g -` A \<inter> space ?B"
+    have "Pi\<^isub>E I F \<subseteq> space (Pi\<^isub>M I M)" "Pi\<^isub>E J F \<subseteq> space (Pi\<^isub>M J M)"
+      using F[rule_format, THEN sets_into_space] by (auto simp: space_PiM)
+    then have X: "?X = (Pi\<^isub>E I F \<times> Pi\<^isub>E J F)"
+      unfolding A_eq by (subst merge_vimage) (auto simp: space_pair_measure space_PiM)
+    have "emeasure ?D A = emeasure ?B ?X"
+      using A by (intro emeasure_distr measurable_merge) (auto simp: sets_PiM)
+    also have "emeasure ?B ?X = (\<Prod>i\<in>I. emeasure (M i) (F i)) * (\<Prod>i\<in>J. emeasure (M i) (F i))"
+      using `finite J` `finite I` F X
+      by (simp add: P.emeasure_pair_measure_Times I.measure_times J.measure_times Pi_iff)
+    also have "\<dots> = (\<Prod>i\<in>I \<union> J. emeasure (M i) (F i))"
       using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod_Un_one)
-    also have "\<dots> = IJ.\<mu> A"
+    also have "\<dots> = emeasure ?P (Pi\<^isub>E (I \<union> J) F)"
       using `finite J` `finite I` F unfolding A
       by (intro IJ.measure_times[symmetric]) auto
-    finally show "IJ.\<mu> A = P.\<mu> (?X A)" by (rule sym)
-  qed (rule measurable_merge[OF IJ])
+    finally show "emeasure ?P A = emeasure ?D A" using A_eq by simp
+  qed
 qed
 
-lemma (in product_sigma_finite) measure_preserving_merge:
-  assumes IJ: "I \<inter> J = {}" and "finite I" "finite J"
-  shows "(\<lambda>(x,y). merge I x J y) \<in> measure_preserving (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)"
-  by (intro measure_preservingI measurable_merge[OF IJ] measure_fold[symmetric] assms)
-
 lemma (in product_sigma_finite) product_positive_integral_fold:
-  assumes IJ[simp]: "I \<inter> J = {}" "finite I" "finite J"
+  assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
   and f: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
   shows "integral\<^isup>P (Pi\<^isub>M (I \<union> J) M) f =
     (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I x J y) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))"
@@ -787,42 +808,38 @@
   interpret I: finite_product_sigma_finite M I by default fact
   interpret J: finite_product_sigma_finite M J by default fact
   interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
-  interpret IJ: finite_product_sigma_finite M "I \<union> J" by default simp
-  have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable P.P"
+  have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
     using measurable_comp[OF measurable_merge[OF IJ(1)] f] by (simp add: comp_def)
   show ?thesis
-    unfolding P.positive_integral_fst_measurable[OF P_borel, simplified]
-  proof (rule P.positive_integral_vimage)
-    show "sigma_algebra IJ.P" by default
-    show "(\<lambda>(x, y). merge I x J y) \<in> measure_preserving P.P IJ.P"
-      using IJ by (rule measure_preserving_merge)
-    show "f \<in> borel_measurable IJ.P" using f by simp
-  qed
+    apply (subst distr_merge[OF IJ, symmetric])
+    apply (subst positive_integral_distr[OF measurable_merge f, OF IJ(1)])
+    apply (subst P.positive_integral_fst_measurable(2)[symmetric, OF P_borel])
+    apply simp
+    done
 qed
 
-lemma (in product_sigma_finite) measure_preserving_component_singelton:
-  "(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)"
-proof (intro measure_preservingI measurable_component_singleton)
+lemma (in product_sigma_finite) distr_singleton:
+  "distr (Pi\<^isub>M {i} M) (M i) (\<lambda>x. x i) = M i" (is "?D = _")
+proof (intro measure_eqI[symmetric])
   interpret I: finite_product_sigma_finite M "{i}" by default simp
-  fix A let ?P = "(\<lambda>x. x i) -` A \<inter> space I.P"
-  assume A: "A \<in> sets (M i)"
-  then have *: "?P = {i} \<rightarrow>\<^isub>E A" using sets_into_space by auto
-  show "I.\<mu> ?P = M.\<mu> i A" unfolding *
-    using A I.measure_times[of "\<lambda>_. A"] by auto
-qed auto
+  fix A assume A: "A \<in> sets (M i)"
+  moreover then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M {i} M) = (\<Pi>\<^isub>E i\<in>{i}. A)"
+    using sets_into_space by (auto simp: space_PiM)
+  ultimately show "emeasure (M i) A = emeasure ?D A"
+    using A I.measure_times[of "\<lambda>_. A"]
+    by (simp add: emeasure_distr measurable_component_singleton)
+qed simp
 
 lemma (in product_sigma_finite) product_positive_integral_singleton:
   assumes f: "f \<in> borel_measurable (M i)"
   shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\<lambda>x. f (x i)) = integral\<^isup>P (M i) f"
 proof -
   interpret I: finite_product_sigma_finite M "{i}" by default simp
-  show ?thesis
-  proof (rule I.positive_integral_vimage[symmetric])
-    show "sigma_algebra (M i)" by (rule sigma_algebras)
-    show "(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)"
-      by (rule measure_preserving_component_singelton)
-    show "f \<in> borel_measurable (M i)" by fact
-  qed
+  from f show ?thesis
+    apply (subst distr_singleton[symmetric])
+    apply (subst positive_integral_distr[OF measurable_component_singleton])
+    apply simp_all
+    done
 qed
 
 lemma (in product_sigma_finite) product_positive_integral_insert:
@@ -832,19 +849,18 @@
 proof -
   interpret I: finite_product_sigma_finite M I by default auto
   interpret i: finite_product_sigma_finite M "{i}" by default auto
-  interpret P: pair_sigma_algebra I.P i.P ..
   have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I"
     using f by auto
   show ?thesis
     unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f]
-  proof (rule I.positive_integral_cong, subst product_positive_integral_singleton)
-    fix x assume x: "x \<in> space I.P"
+  proof (rule positive_integral_cong, subst product_positive_integral_singleton)
+    fix x assume x: "x \<in> space (Pi\<^isub>M I M)"
     let ?f = "\<lambda>y. f (restrict (x(i := y)) (insert i I))"
     have f'_eq: "\<And>y. ?f y = f (x(i := y))"
-      using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def)
+      using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def space_PiM)
     show "?f \<in> borel_measurable (M i)" unfolding f'_eq
-      using measurable_comp[OF measurable_component_update f] x `i \<notin> I`
-      by (simp add: comp_def)
+      using measurable_comp[OF measurable_component_update f, OF x `i \<notin> I`]
+      unfolding comp_def .
     show "integral\<^isup>P (M i) ?f = \<integral>\<^isup>+ y. f (x(i:=y)) \<partial>M i"
       unfolding f'_eq by simp
   qed
@@ -856,10 +872,6 @@
   and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x"
   shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))"
 using assms proof induct
-  case empty
-  interpret finite_product_sigma_finite M "{}" by default auto
-  show ?case by simp
-next
   case (insert i I)
   note `finite I`[intro, simp]
   interpret I: finite_product_sigma_finite M I by default auto
@@ -867,16 +879,16 @@
     using insert by (auto intro!: setprod_cong)
   have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)"
     using sets_into_space insert
-    by (intro sigma_algebra.borel_measurable_ereal_setprod sigma_algebra_product_algebra
+    by (intro borel_measurable_ereal_setprod
               measurable_comp[OF measurable_component_singleton, unfolded comp_def])
        auto
   then show ?case
     apply (simp add: product_positive_integral_insert[OF insert(1,2) prod])
-    apply (simp add: insert * pos borel setprod_ereal_pos M.positive_integral_multc)
-    apply (subst I.positive_integral_cmult)
-    apply (auto simp add: pos borel insert setprod_ereal_pos positive_integral_positive)
+    apply (simp add: insert(2-) * pos borel setprod_ereal_pos positive_integral_multc)
+    apply (subst positive_integral_cmult)
+    apply (auto simp add: pos borel insert(2-) setprod_ereal_pos positive_integral_positive)
     done
-qed
+qed (simp add: space_PiM)
 
 lemma (in product_sigma_finite) product_integral_singleton:
   assumes f: "f \<in> borel_measurable (M i)"
@@ -899,48 +911,44 @@
   interpret J: finite_product_sigma_finite M J by default fact
   have "finite (I \<union> J)" using fin by auto
   interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact
-  interpret P: pair_sigma_finite I.P J.P by default
+  interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default
   let ?M = "\<lambda>(x, y). merge I x J y"
   let ?f = "\<lambda>x. f (?M x)"
+  from f have f_borel: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
+    by auto
+  have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
+    using measurable_comp[OF measurable_merge[OF IJ(1)] f_borel] by (simp add: comp_def)
+  have f_int: "integrable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ?f"
+    by (rule integrable_distr[OF measurable_merge[OF IJ]]) (simp add: distr_merge[OF IJ fin] f)
   show ?thesis
-  proof (subst P.integrable_fst_measurable(2)[of ?f, simplified])
-    have 1: "sigma_algebra IJ.P" by default
-    have 2: "?M \<in> measure_preserving P.P IJ.P" using measure_preserving_merge[OF assms(1,2,3)] .
-    have 3: "integrable (Pi\<^isub>M (I \<union> J) M) f" by fact
-    then have 4: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)"
-      by (simp add: integrable_def)
-    show "integrable P.P ?f"
-      by (rule P.integrable_vimage[where f=f, OF 1 2 3])
-    show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f"
-      by (rule P.integral_vimage[where f=f, OF 1 2 4])
-  qed
+    apply (subst distr_merge[symmetric, OF IJ fin])
+    apply (subst integral_distr[OF measurable_merge[OF IJ] f_borel])
+    apply (subst P.integrable_fst_measurable(2)[symmetric, OF f_int])
+    apply simp
+    done
 qed
 
 lemma (in product_sigma_finite) product_integral_insert:
-  assumes [simp]: "finite I" "i \<notin> I"
+  assumes I: "finite I" "i \<notin> I"
     and f: "integrable (Pi\<^isub>M (insert i I) M) f"
   shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^isub>M I M)"
 proof -
-  interpret I: finite_product_sigma_finite M I by default auto
-  interpret I': finite_product_sigma_finite M "insert i I" by default auto
-  interpret i: finite_product_sigma_finite M "{i}" by default auto
-  interpret P: pair_sigma_finite I.P i.P ..
-  have IJ: "I \<inter> {i} = {}" by auto
-  show ?thesis
-    unfolding product_integral_fold[OF IJ, simplified, OF f]
-  proof (rule I.integral_cong, subst product_integral_singleton)
-    fix x assume x: "x \<in> space I.P"
-    let ?f = "\<lambda>y. f (restrict (x(i := y)) (insert i I))"
-    have f'_eq: "\<And>y. ?f y = f (x(i := y))"
-      using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def)
-    have f: "f \<in> borel_measurable I'.P" using f unfolding integrable_def by auto
-    show "?f \<in> borel_measurable (M i)"
-      unfolding measurable_cong[OF f'_eq]
-      using measurable_comp[OF measurable_component_update f] x `i \<notin> I`
-      by (simp add: comp_def)
-    show "integral\<^isup>L (M i) ?f = integral\<^isup>L (M i) (\<lambda>y. f (x(i := y)))"
-      unfolding f'_eq by simp
+  have "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = integral\<^isup>L (Pi\<^isub>M (I \<union> {i}) M) f"
+    by simp
+  also have "\<dots> = (\<integral>x. (\<integral>y. f (merge I x {i} y) \<partial>Pi\<^isub>M {i} M) \<partial>Pi\<^isub>M I M)"
+    using f I by (intro product_integral_fold) auto
+  also have "\<dots> = (\<integral>x. (\<integral>y. f (x(i := y)) \<partial>M i) \<partial>Pi\<^isub>M I M)"
+  proof (rule integral_cong, subst product_integral_singleton[symmetric])
+    fix x assume x: "x \<in> space (Pi\<^isub>M I M)"
+    have f_borel: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)"
+      using f by auto
+    show "(\<lambda>y. f (x(i := y))) \<in> borel_measurable (M i)"
+      using measurable_comp[OF measurable_component_update f_borel, OF x `i \<notin> I`]
+      unfolding comp_def .
+    from x I show "(\<integral> y. f (merge I x {i} y) \<partial>Pi\<^isub>M {i} M) = (\<integral> xa. f (x(i := xa i)) \<partial>Pi\<^isub>M {i} M)"
+      by (auto intro!: integral_cong arg_cong[where f=f] simp: merge_def space_PiM extensional_def)
   qed
+  finally show ?thesis .
 qed
 
 lemma (in product_sigma_finite) product_integrable_setprod:
@@ -951,24 +959,23 @@
   interpret finite_product_sigma_finite M I by default fact
   have f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)"
     using integrable unfolding integrable_def by auto
-  then have borel: "?f \<in> borel_measurable P"
-    using measurable_comp[OF measurable_component_singleton f]
-    by (auto intro!: borel_measurable_setprod simp: comp_def)
+  have borel: "?f \<in> borel_measurable (Pi\<^isub>M I M)"
+    using measurable_comp[OF measurable_component_singleton[of _ I M] f] by (auto simp: comp_def)
   moreover have "integrable (Pi\<^isub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)"
   proof (unfold integrable_def, intro conjI)
-    show "(\<lambda>x. abs (?f x)) \<in> borel_measurable P"
+    show "(\<lambda>x. abs (?f x)) \<in> borel_measurable (Pi\<^isub>M I M)"
       using borel by auto
-    have "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. ereal (abs (f i (x i)))) \<partial>P)"
+    have "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>Pi\<^isub>M I M) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. ereal (abs (f i (x i)))) \<partial>Pi\<^isub>M I M)"
       by (simp add: setprod_ereal abs_setprod)
     also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. ereal (abs (f i x)) \<partial>M i))"
       using f by (subst product_positive_integral_setprod) auto
     also have "\<dots> < \<infinity>"
-      using integrable[THEN M.integrable_abs]
-      by (simp add: setprod_PInf integrable_def M.positive_integral_positive)
-    finally show "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto
-    have "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)"
+      using integrable[THEN integrable_abs]
+      by (simp add: setprod_PInf integrable_def positive_integral_positive)
+    finally show "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>(Pi\<^isub>M I M)) \<noteq> \<infinity>" by auto
+    have "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^isub>M I M)) = (\<integral>\<^isup>+x. 0 \<partial>(Pi\<^isub>M I M))"
       by (intro positive_integral_cong_pos) auto
-    then show "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp
+    then show "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>(Pi\<^isub>M I M)) \<noteq> \<infinity>" by simp
   qed
   ultimately show ?thesis
     by (rule integrable_abs_iff[THEN iffD1])
@@ -992,7 +999,69 @@
     using `i \<notin> I` by (auto intro!: setprod_cong)
   show ?case
     unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]]
-    by (simp add: * insert integral_multc I.integral_cmult[OF prod] subset_insertI)
+    by (simp add: * insert integral_multc integral_cmult[OF prod] subset_insertI)
 qed
 
-end
\ No newline at end of file
+lemma sigma_prod_algebra_sigma_eq:
+  fixes E :: "'i \<Rightarrow> 'a set set"
+  assumes "finite I"
+  assumes S_mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)"
+    and S_union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (M i)"
+    and S_in_E: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> E i"
+  assumes E_closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M i))"
+    and E_generates: "\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sigma_sets (space (M i)) (E i)"
+  defines "P == { Pi\<^isub>E I F | F. \<forall>i\<in>I. F i \<in> E i }"
+  shows "sets (PiM I M) = sigma_sets (space (PiM I M)) P"
+proof
+  let ?P = "sigma (space (Pi\<^isub>M I M)) P"
+  have P_closed: "P \<subseteq> Pow (space (Pi\<^isub>M I M))"
+    using E_closed by (auto simp: space_PiM P_def Pi_iff subset_eq)
+  then have space_P: "space ?P = (\<Pi>\<^isub>E i\<in>I. space (M i))"
+    by (simp add: space_PiM)
+  have "sets (PiM I M) =
+      sigma_sets (space ?P) {{f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A} |i A. i \<in> I \<and> A \<in> sets (M i)}"
+    using sets_PiM_single[of I M] by (simp add: space_P)
+  also have "\<dots> \<subseteq> sets (sigma (space (PiM I M)) P)"
+  proof (safe intro!: sigma_sets_subset)
+    fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
+    have "(\<lambda>x. x i) \<in> measurable ?P (sigma (space (M i)) (E i))"
+    proof (subst measurable_iff_measure_of)
+      show "E i \<subseteq> Pow (space (M i))" using `i \<in> I` by fact
+      from space_P `i \<in> I` show "(\<lambda>x. x i) \<in> space ?P \<rightarrow> space (M i)"
+        by (auto simp: Pi_iff)
+      show "\<forall>A\<in>E i. (\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
+      proof
+        fix A assume A: "A \<in> E i"
+        then have "(\<lambda>x. x i) -` A \<inter> space ?P = (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))"
+          using E_closed `i \<in> I` by (auto simp: space_P Pi_iff subset_eq split: split_if_asm)
+        also have "\<dots> = (\<Pi>\<^isub>E j\<in>I. \<Union>n. if i = j then A else S j n)"
+          by (intro PiE_cong) (simp add: S_union)
+        also have "\<dots> = (\<Union>n. \<Pi>\<^isub>E j\<in>I. if i = j then A else S j n)"
+          using S_mono
+          by (subst Pi_UN[symmetric, OF `finite I`]) (auto simp: incseq_def)
+        also have "\<dots> \<in> sets ?P"
+        proof (safe intro!: countable_UN)
+          fix n show "(\<Pi>\<^isub>E j\<in>I. if i = j then A else S j n) \<in> sets ?P"
+            using A S_in_E
+            by (simp add: P_closed)
+               (auto simp: P_def subset_eq intro!: exI[of _ "\<lambda>j. if i = j then A else S j n"])
+        qed
+        finally show "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
+          using P_closed by simp
+      qed
+    qed
+    from measurable_sets[OF this, of A] A `i \<in> I` E_closed
+    have "(\<lambda>x. x i) -` A \<inter> space ?P \<in> sets ?P"
+      by (simp add: E_generates)
+    also have "(\<lambda>x. x i) -` A \<inter> space ?P = {f \<in> \<Pi>\<^isub>E i\<in>I. space (M i). f i \<in> A}"
+      using P_closed by (auto simp: space_PiM)
+    finally show "\<dots> \<in> sets ?P" .
+  qed
+  finally show "sets (PiM I M) \<subseteq> sigma_sets (space (PiM I M)) P"
+    by (simp add: P_closed)
+  show "sigma_sets (space (PiM I M)) P \<subseteq> sets (PiM I M)"
+    using `finite I`
+    by (auto intro!: sigma_sets_subset simp: E_generates P_def)
+qed
+
+end
--- a/src/HOL/Probability/Independent_Family.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Independent_Family.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -5,7 +5,7 @@
 header {* Independent families of events, event sets, and random variables *}
 
 theory Independent_Family
-  imports Probability_Measure
+  imports Probability_Measure Infinite_Product_Measure
 begin
 
 lemma INT_decseq_offset:
@@ -44,7 +44,7 @@
 definition (in prob_space)
   "indep_var Ma A Mb B \<longleftrightarrow> indep_vars (bool_case Ma Mb) (bool_case A B) UNIV"
 
-lemma (in prob_space) indep_sets_cong[cong]:
+lemma (in prob_space) indep_sets_cong:
   "I = J \<Longrightarrow> (\<And>i. i \<in> I \<Longrightarrow> F i = G i) \<Longrightarrow> indep_sets F I \<longleftrightarrow> indep_sets G J"
   by (simp add: indep_sets_def, intro conj_cong all_cong imp_cong ball_cong) blast+
 
@@ -135,7 +135,7 @@
 
 lemma (in prob_space) indep_sets_dynkin:
   assumes indep: "indep_sets F I"
-  shows "indep_sets (\<lambda>i. sets (dynkin \<lparr> space = space M, sets = F i \<rparr>)) I"
+  shows "indep_sets (\<lambda>i. dynkin (space M) (F i)) I"
     (is "indep_sets ?F I")
 proof (subst indep_sets_finite_index_sets, intro allI impI ballI)
   fix J assume "finite J" "J \<subseteq> I" "J \<noteq> {}"
@@ -193,7 +193,7 @@
           qed
         qed }
       note indep_sets_insert = this
-      have "dynkin_system \<lparr> space = space M, sets = ?D \<rparr>"
+      have "dynkin_system (space M) ?D"
       proof (rule dynkin_systemI', simp_all cong del: indep_sets_cong, safe)
         show "indep_sets (G(j := {{}})) K"
           by (rule indep_sets_insert) auto
@@ -206,7 +206,7 @@
             using G by auto
           have "prob ((\<Inter>j\<in>J. A j) \<inter> (space M - X)) =
               prob ((\<Inter>j\<in>J. A j) - (\<Inter>i\<in>insert j J. (A(j := X)) i))"
-            using A_sets sets_into_space X `J \<noteq> {}`
+            using A_sets sets_into_space[of _ M] X `J \<noteq> {}`
             by (auto intro!: arg_cong[where f=prob] split: split_if_asm)
           also have "\<dots> = prob (\<Inter>j\<in>J. A j) - prob (\<Inter>i\<in>insert j J. (A(j := X)) i)"
             using J `J \<noteq> {}` `j \<notin> J` A_sets X sets_into_space
@@ -264,9 +264,8 @@
             by (auto dest!: sums_unique)
         qed (insert F, auto)
       qed (insert sets_into_space, auto)
-      then have mono: "sets (dynkin \<lparr>space = space M, sets = G j\<rparr>) \<subseteq>
-        sets \<lparr>space = space M, sets = {E \<in> events. indep_sets (G(j := {E})) K}\<rparr>"
-      proof (rule dynkin_system.dynkin_subset, simp_all cong del: indep_sets_cong, safe)
+      then have mono: "dynkin (space M) (G j) \<subseteq> {E \<in> events. indep_sets (G(j := {E})) K}"
+      proof (rule dynkin_system.dynkin_subset, safe)
         fix X assume "X \<in> G j"
         then show "X \<in> events" using G `j \<in> K` by auto
         from `indep_sets G K`
@@ -292,20 +291,20 @@
             by (intro indep_setsD[OF G(1)]) auto
         qed
       qed
-      then have "indep_sets (G(j:=sets (dynkin \<lparr>space = space M, sets = G j\<rparr>))) K"
+      then have "indep_sets (G(j := dynkin (space M) (G j))) K"
         by (rule indep_sets_mono_sets) (insert mono, auto)
       then show ?case
         by (rule indep_sets_mono_sets) (insert `j \<in> K` `j \<notin> J`, auto simp: G_def)
     qed (insert `indep_sets F K`, simp) }
   from this[OF `indep_sets F J` `finite J` subset_refl]
-  show "indep_sets (\<lambda>i. sets (dynkin \<lparr> space = space M, sets = F i \<rparr>)) J"
+  show "indep_sets ?F J"
     by (rule indep_sets_mono_sets) auto
 qed
 
 lemma (in prob_space) indep_sets_sigma:
   assumes indep: "indep_sets F I"
-  assumes stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable \<lparr> space = space M, sets = F i \<rparr>"
-  shows "indep_sets (\<lambda>i. sets (sigma \<lparr> space = space M, sets = F i \<rparr>)) I"
+  assumes stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable (F i)"
+  shows "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I"
 proof -
   from indep_sets_dynkin[OF indep]
   show ?thesis
@@ -316,18 +315,12 @@
   qed
 qed
 
-lemma (in prob_space) indep_sets_sigma_sets:
-  assumes "indep_sets F I"
-  assumes "\<And>i. i \<in> I \<Longrightarrow> Int_stable \<lparr> space = space M, sets = F i \<rparr>"
-  shows "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I"
-  using indep_sets_sigma[OF assms] by (simp add: sets_sigma)
-
 lemma (in prob_space) indep_sets_sigma_sets_iff:
-  assumes "\<And>i. i \<in> I \<Longrightarrow> Int_stable \<lparr> space = space M, sets = F i \<rparr>"
+  assumes "\<And>i. i \<in> I \<Longrightarrow> Int_stable (F i)"
   shows "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I \<longleftrightarrow> indep_sets F I"
 proof
   assume "indep_sets F I" then show "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I"
-    by (rule indep_sets_sigma_sets) fact
+    by (rule indep_sets_sigma) fact
 next
   assume "indep_sets (\<lambda>i. sigma_sets (space M) (F i)) I" then show "indep_sets F I"
     by (rule indep_sets_mono_sets) (intro subsetI sigma_sets.Basic)
@@ -361,15 +354,14 @@
 
 lemma (in prob_space) indep_set_sigma_sets:
   assumes "indep_set A B"
-  assumes A: "Int_stable \<lparr> space = space M, sets = A \<rparr>"
-  assumes B: "Int_stable \<lparr> space = space M, sets = B \<rparr>"
+  assumes A: "Int_stable A" and B: "Int_stable B"
   shows "indep_set (sigma_sets (space M) A) (sigma_sets (space M) B)"
 proof -
   have "indep_sets (\<lambda>i. sigma_sets (space M) (case i of True \<Rightarrow> A | False \<Rightarrow> B)) UNIV"
-  proof (rule indep_sets_sigma_sets)
+  proof (rule indep_sets_sigma)
     show "indep_sets (bool_case A B) UNIV"
       by (rule `indep_set A B`[unfolded indep_set_def])
-    fix i show "Int_stable \<lparr>space = space M, sets = case i of True \<Rightarrow> A | False \<Rightarrow> B\<rparr>"
+    fix i show "Int_stable (case i of True \<Rightarrow> A | False \<Rightarrow> B)"
       using A B by (cases i) auto
   qed
   then show ?thesis
@@ -380,7 +372,7 @@
 lemma (in prob_space) indep_sets_collect_sigma:
   fixes I :: "'j \<Rightarrow> 'i set" and J :: "'j set" and E :: "'i \<Rightarrow> 'a set set"
   assumes indep: "indep_sets E (\<Union>j\<in>J. I j)"
-  assumes Int_stable: "\<And>i j. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> Int_stable \<lparr>space = space M, sets = E i\<rparr>"
+  assumes Int_stable: "\<And>i j. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> Int_stable (E i)"
   assumes disjoint: "disjoint_family_on I J"
   shows "indep_sets (\<lambda>j. sigma_sets (space M) (\<Union>i\<in>I j. E i)) J"
 proof -
@@ -389,30 +381,29 @@
   from indep have E: "\<And>j i. j \<in> J \<Longrightarrow> i \<in> I j \<Longrightarrow> E i \<subseteq> events"
     unfolding indep_sets_def by auto
   { fix j
-    let ?S = "sigma \<lparr> space = space M, sets = (\<Union>i\<in>I j. E i) \<rparr>"
+    let ?S = "sigma_sets (space M) (\<Union>i\<in>I j. E i)"
     assume "j \<in> J"
-    from E[OF this] interpret S: sigma_algebra ?S
-      using sets_into_space by (intro sigma_algebra_sigma) auto
+    from E[OF this] interpret S: sigma_algebra "space M" ?S
+      using sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto
 
     have "sigma_sets (space M) (\<Union>i\<in>I j. E i) = sigma_sets (space M) (?E j)"
     proof (rule sigma_sets_eqI)
       fix A assume "A \<in> (\<Union>i\<in>I j. E i)"
       then guess i ..
       then show "A \<in> sigma_sets (space M) (?E j)"
-        by (auto intro!: sigma_sets.intros exI[of _ "{i}"] exI[of _ "\<lambda>i. A"])
+        by (auto intro!: sigma_sets.intros(2-) exI[of _ "{i}"] exI[of _ "\<lambda>i. A"])
     next
       fix A assume "A \<in> ?E j"
       then obtain E' K where "finite K" "K \<noteq> {}" "K \<subseteq> I j" "\<And>k. k \<in> K \<Longrightarrow> E' k \<in> E k"
         and A: "A = (\<Inter>k\<in>K. E' k)"
         by auto
-      then have "A \<in> sets ?S" unfolding A
-        by (safe intro!: S.finite_INT)
-           (auto simp: sets_sigma intro!: sigma_sets.Basic)
+      then have "A \<in> ?S" unfolding A
+        by (safe intro!: S.finite_INT) auto
       then show "A \<in> sigma_sets (space M) (\<Union>i\<in>I j. E i)"
-        by (simp add: sets_sigma)
+        by simp
     qed }
   moreover have "indep_sets (\<lambda>j. sigma_sets (space M) (?E j)) J"
-  proof (rule indep_sets_sigma_sets)
+  proof (rule indep_sets_sigma)
     show "indep_sets ?E J"
     proof (intro indep_setsI)
       fix j assume "j \<in> J" with E show "?E j \<subseteq> events" by (force  intro!: finite_INT)
@@ -460,7 +451,7 @@
     qed
   next
     fix j assume "j \<in> J"
-    show "Int_stable \<lparr> space = space M, sets = ?E j \<rparr>"
+    show "Int_stable (?E j)"
     proof (rule Int_stableI)
       fix a assume "a \<in> ?E j" then obtain Ka Ea
         where a: "a = (\<Inter>k\<in>Ka. Ea k)" "finite Ka" "Ka \<noteq> {}" "Ka \<subseteq> I j" "\<And>k. k\<in>Ka \<Longrightarrow> Ea k \<in> E k" by auto
@@ -482,12 +473,12 @@
 
 lemma (in prob_space) terminal_events_sets:
   assumes A: "\<And>i. A i \<subseteq> events"
-  assumes "\<And>i::nat. sigma_algebra \<lparr>space = space M, sets = A i\<rparr>"
+  assumes "\<And>i::nat. sigma_algebra (space M) (A i)"
   assumes X: "X \<in> terminal_events A"
   shows "X \<in> events"
 proof -
   let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
-  interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
+  interpret A: sigma_algebra "space M" "A i" for i by fact
   from X have "\<And>n. X \<in> sigma_sets (space M) (UNION {n..} A)" by (auto simp: terminal_events_def)
   from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp
   then show "X \<in> events"
@@ -495,12 +486,12 @@
 qed
 
 lemma (in prob_space) sigma_algebra_terminal_events:
-  assumes "\<And>i::nat. sigma_algebra \<lparr>space = space M, sets = A i\<rparr>"
-  shows "sigma_algebra \<lparr> space = space M, sets = terminal_events A \<rparr>"
+  assumes "\<And>i::nat. sigma_algebra (space M) (A i)"
+  shows "sigma_algebra (space M) (terminal_events A)"
   unfolding terminal_events_def
 proof (simp add: sigma_algebra_iff2, safe)
   let ?A = "(\<Inter>n. sigma_sets (space M) (UNION {n..} A))"
-  interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
+  interpret A: sigma_algebra "space M" "A i" for i by fact
   { fix X x assume "X \<in> ?A" "x \<in> X"
     then have "\<And>n. X \<in> sigma_sets (space M) (UNION {n..} A)" by auto
     from this[of 0] have "X \<in> sigma_sets (space M) (UNION UNIV A)" by simp
@@ -515,29 +506,29 @@
 lemma (in prob_space) kolmogorov_0_1_law:
   fixes A :: "nat \<Rightarrow> 'a set set"
   assumes A: "\<And>i. A i \<subseteq> events"
-  assumes "\<And>i::nat. sigma_algebra \<lparr>space = space M, sets = A i\<rparr>"
+  assumes "\<And>i::nat. sigma_algebra (space M) (A i)"
   assumes indep: "indep_sets A UNIV"
   and X: "X \<in> terminal_events A"
   shows "prob X = 0 \<or> prob X = 1"
 proof -
-  let ?D = "\<lparr> space = space M, sets = {D \<in> events. prob (X \<inter> D) = prob X * prob D} \<rparr>"
-  interpret A: sigma_algebra "\<lparr>space = space M, sets = A i\<rparr>" for i by fact
-  interpret T: sigma_algebra "\<lparr> space = space M, sets = terminal_events A \<rparr>"
+  let ?D = "{D \<in> events. prob (X \<inter> D) = prob X * prob D}"
+  interpret A: sigma_algebra "space M" "A i" for i by fact
+  interpret T: sigma_algebra "space M" "terminal_events A"
     by (rule sigma_algebra_terminal_events) fact
   have "X \<subseteq> space M" using T.space_closed X by auto
 
   have X_in: "X \<in> events"
     by (rule terminal_events_sets) fact+
 
-  interpret D: dynkin_system ?D
+  interpret D: dynkin_system "space M" ?D
   proof (rule dynkin_systemI)
-    fix D assume "D \<in> sets ?D" then show "D \<subseteq> space ?D"
+    fix D assume "D \<in> ?D" then show "D \<subseteq> space M"
       using sets_into_space by auto
   next
-    show "space ?D \<in> sets ?D"
+    show "space M \<in> ?D"
       using prob_space `X \<subseteq> space M` by (simp add: Int_absorb2)
   next
-    fix A assume A: "A \<in> sets ?D"
+    fix A assume A: "A \<in> ?D"
     have "prob (X \<inter> (space M - A)) = prob (X - (X \<inter> A))"
       using `X \<subseteq> space M` by (auto intro!: arg_cong[where f=prob])
     also have "\<dots> = prob X - prob (X \<inter> A)"
@@ -547,10 +538,10 @@
     also have "\<dots> = prob X * prob (space M - A)"
       using X_in A sets_into_space
       by (subst finite_measure_Diff) (auto simp: field_simps)
-    finally show "space ?D - A \<in> sets ?D"
+    finally show "space M - A \<in> ?D"
       using A `X \<subseteq> space M` by auto
   next
-    fix F :: "nat \<Rightarrow> 'a set" assume dis: "disjoint_family F" and "range F \<subseteq> sets ?D"
+    fix F :: "nat \<Rightarrow> 'a set" assume dis: "disjoint_family F" and "range F \<subseteq> ?D"
     then have F: "range F \<subseteq> events" "\<And>i. prob (X \<inter> F i) = prob X * prob (F i)"
       by auto
     have "(\<lambda>i. prob (X \<inter> F i)) sums prob (\<Union>i. X \<inter> F i)"
@@ -566,7 +557,7 @@
       by (intro sums_mult finite_measure_UNION F dis)
     ultimately have "prob (X \<inter> (\<Union>i. F i)) = prob X * prob (\<Union>i. F i)"
       by (auto dest!: sums_unique)
-    with F show "(\<Union>i. F i) \<in> sets ?D"
+    with F show "(\<Union>i. F i) \<in> ?D"
       by auto
   qed
 
@@ -579,7 +570,7 @@
       show "disjoint_family (bool_case {..n} {Suc n..})"
         unfolding disjoint_family_on_def by (auto split: bool.split)
       fix m
-      show "Int_stable \<lparr>space = space M, sets = A m\<rparr>"
+      show "Int_stable (A m)"
         unfolding Int_stable_def using A.Int by auto
     qed
     also have "(\<lambda>b. sigma_sets (space M) (\<Union>m\<in>bool_case {..n} {Suc n..} b. A m)) =
@@ -588,7 +579,7 @@
     finally have indep: "indep_set (sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) (sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m))"
       unfolding indep_set_def by simp
 
-    have "sigma_sets (space M) (\<Union>m\<in>{..n}. A m) \<subseteq> sets ?D"
+    have "sigma_sets (space M) (\<Union>m\<in>{..n}. A m) \<subseteq> ?D"
     proof (simp add: subset_eq, rule)
       fix D assume D: "D \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
       have "X \<in> sigma_sets (space M) (\<Union>m\<in>{Suc n..}. A m)"
@@ -597,22 +588,27 @@
       show "D \<in> events \<and> prob (X \<inter> D) = prob X * prob D"
         by (auto simp add: ac_simps)
     qed }
-  then have "(\<Union>n. sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) \<subseteq> sets ?D" (is "?A \<subseteq> _")
+  then have "(\<Union>n. sigma_sets (space M) (\<Union>m\<in>{..n}. A m)) \<subseteq> ?D" (is "?A \<subseteq> _")
     by auto
 
-  have "sigma \<lparr> space = space M, sets = ?A \<rparr> =
-    dynkin \<lparr> space = space M, sets = ?A \<rparr>" (is "sigma ?UA = dynkin ?UA")
+  note `X \<in> terminal_events A`
+  also {
+    have "\<And>n. sigma_sets (space M) (\<Union>i\<in>{n..}. A i) \<subseteq> sigma_sets (space M) ?A"
+      by (intro sigma_sets_subseteq UN_mono) auto
+   then have "terminal_events A \<subseteq> sigma_sets (space M) ?A"
+      unfolding terminal_events_def by auto }
+  also have "sigma_sets (space M) ?A = dynkin (space M) ?A"
   proof (rule sigma_eq_dynkin)
     { fix B n assume "B \<in> sigma_sets (space M) (\<Union>m\<in>{..n}. A m)"
       then have "B \<subseteq> space M"
-        by induct (insert A sets_into_space, auto) }
-    then show "sets ?UA \<subseteq> Pow (space ?UA)" by auto
-    show "Int_stable ?UA"
+        by induct (insert A sets_into_space[of _ M], auto) }
+    then show "?A \<subseteq> Pow (space M)" by auto
+    show "Int_stable ?A"
     proof (rule Int_stableI)
       fix a assume "a \<in> ?A" then guess n .. note a = this
       fix b assume "b \<in> ?A" then guess m .. note b = this
-      interpret Amn: sigma_algebra "sigma \<lparr>space = space M, sets = (\<Union>i\<in>{..max m n}. A i)\<rparr>"
-        using A sets_into_space by (intro sigma_algebra_sigma) auto
+      interpret Amn: sigma_algebra "space M" "sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
+        using A sets_into_space[of _ M] by (intro sigma_algebra_sigma_sets) auto
       have "sigma_sets (space M) (\<Union>i\<in>{..n}. A i) \<subseteq> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
         by (intro sigma_sets_subseteq UN_mono) auto
       with a have "a \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto
@@ -621,23 +617,13 @@
         by (intro sigma_sets_subseteq UN_mono) auto
       with b have "b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)" by auto
       ultimately have "a \<inter> b \<in> sigma_sets (space M) (\<Union>i\<in>{..max m n}. A i)"
-        using Amn.Int[of a b] by (simp add: sets_sigma)
+        using Amn.Int[of a b] by simp
       then show "a \<inter> b \<in> (\<Union>n. sigma_sets (space M) (\<Union>i\<in>{..n}. A i))" by auto
     qed
   qed
-  moreover have "sets (dynkin ?UA) \<subseteq> sets ?D"
-  proof (rule D.dynkin_subset)
-    show "sets ?UA \<subseteq> sets ?D" using `?A \<subseteq> sets ?D` by auto
-  qed simp
-  ultimately have "sets (sigma ?UA) \<subseteq> sets ?D" by simp
-  moreover
-  have "\<And>n. sigma_sets (space M) (\<Union>i\<in>{n..}. A i) \<subseteq> sigma_sets (space M) ?A"
-    by (intro sigma_sets_subseteq UN_mono) (auto intro: sigma_sets.Basic)
-  then have "terminal_events A \<subseteq> sets (sigma ?UA)"
-    unfolding sets_sigma terminal_events_def by auto
-  moreover note `X \<in> terminal_events A`
-  ultimately have "X \<in> sets ?D" by auto
-  then show ?thesis by auto
+  also have "dynkin (space M) ?A \<subseteq> ?D"
+    using `?A \<subseteq> ?D` by (auto intro!: D.dynkin_subset)
+  finally show ?thesis by auto
 qed
 
 lemma (in prob_space) borel_0_1_law:
@@ -648,14 +634,14 @@
   show "\<And>i. sigma_sets (space M) {F i} \<subseteq> events"
     using F(1) sets_into_space
     by (subst sigma_sets_singleton) auto
-  { fix i show "sigma_algebra \<lparr>space = space M, sets = sigma_sets (space M) {F i}\<rparr>"
-      using sigma_algebra_sigma[of "\<lparr>space = space M, sets = {F i}\<rparr>"] F sets_into_space
-      by (auto simp add: sigma_def) }
+  { fix i show "sigma_algebra (space M) (sigma_sets (space M) {F i})"
+      using sigma_algebra_sigma_sets[of "{F i}" "space M"] F sets_into_space
+      by auto }
   show "indep_sets (\<lambda>i. sigma_sets (space M) {F i}) UNIV"
-  proof (rule indep_sets_sigma_sets)
+  proof (rule indep_sets_sigma)
     show "indep_sets (\<lambda>i. {F i}) UNIV"
       unfolding indep_sets_singleton_iff_indep_events by fact
-    fix i show "Int_stable \<lparr>space = space M, sets = {F i}\<rparr>"
+    fix i show "Int_stable {F i}"
       unfolding Int_stable_def by simp
   qed
   let ?Q = "\<lambda>n. \<Union>i\<in>{n..}. F i"
@@ -663,17 +649,17 @@
     unfolding terminal_events_def
   proof
     fix j
-    interpret S: sigma_algebra "sigma \<lparr> space = space M, sets = (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})\<rparr>"
+    interpret S: sigma_algebra "space M" "sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
       using order_trans[OF F(1) space_closed]
-      by (intro sigma_algebra_sigma) (simp add: sigma_sets_singleton subset_eq)
+      by (intro sigma_algebra_sigma_sets) (simp add: sigma_sets_singleton subset_eq)
     have "(\<Inter>n. ?Q n) = (\<Inter>n\<in>{j..}. ?Q n)"
       by (intro decseq_SucI INT_decseq_offset UN_mono) auto
-    also have "\<dots> \<in> sets (sigma \<lparr> space = space M, sets = (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})\<rparr>)"
+    also have "\<dots> \<in> sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
       using order_trans[OF F(1) space_closed]
       by (safe intro!: S.countable_INT S.countable_UN)
-         (auto simp: sets_sigma sigma_sets_singleton intro!: sigma_sets.Basic bexI)
+         (auto simp: sigma_sets_singleton intro!: sigma_sets.Basic bexI)
     finally show "(\<Inter>n. ?Q n) \<in> sigma_sets (space M) (\<Union>i\<in>{j..}. sigma_sets (space M) {F i})"
-      by (simp add: sets_sigma)
+      by simp
   qed
 qed
 
@@ -710,84 +696,84 @@
 lemma (in prob_space) indep_vars_finite:
   fixes I :: "'i set"
   assumes I: "I \<noteq> {}" "finite I"
-    and rv: "\<And>i. i \<in> I \<Longrightarrow> random_variable (sigma (M' i)) (X i)"
-    and Int_stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable (M' i)"
-    and space: "\<And>i. i \<in> I \<Longrightarrow> space (M' i) \<in> sets (M' i)"
-  shows "indep_vars (\<lambda>i. sigma (M' i)) X I \<longleftrightarrow>
-    (\<forall>A\<in>(\<Pi> i\<in>I. sets (M' i)). prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = (\<Prod>j\<in>I. prob (X j -` A j \<inter> space M)))"
+    and M': "\<And>i. i \<in> I \<Longrightarrow> sets (M' i) = sigma_sets (space (M' i)) (E i)"
+    and rv: "\<And>i. i \<in> I \<Longrightarrow> random_variable (M' i) (X i)"
+    and Int_stable: "\<And>i. i \<in> I \<Longrightarrow> Int_stable (E i)"
+    and space: "\<And>i. i \<in> I \<Longrightarrow> space (M' i) \<in> E i" and closed: "\<And>i. i \<in> I \<Longrightarrow> E i \<subseteq> Pow (space (M' i))"
+  shows "indep_vars M' X I \<longleftrightarrow>
+    (\<forall>A\<in>(\<Pi> i\<in>I. E i). prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = (\<Prod>j\<in>I. prob (X j -` A j \<inter> space M)))"
 proof -
   from rv have X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> space M \<rightarrow> space (M' i)"
     unfolding measurable_def by simp
 
   { fix i assume "i\<in>I"
-    have "sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (sigma (M' i))}
-      = sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
-      unfolding sigma_sets_vimage_commute[OF X, OF `i \<in> I`]
+    from closed[OF `i \<in> I`]
+    have "sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}
+      = sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> E i}"
+      unfolding sigma_sets_vimage_commute[OF X, OF `i \<in> I`, symmetric] M'[OF `i \<in> I`]
       by (subst sigma_sets_sigma_sets_eq) auto }
-  note this[simp]
+  note sigma_sets_X = this
 
   { fix i assume "i\<in>I"
-    have "Int_stable \<lparr>space = space M, sets = {X i -` A \<inter> space M |A. A \<in> sets (M' i)}\<rparr>"
+    have "Int_stable {X i -` A \<inter> space M |A. A \<in> E i}"
     proof (rule Int_stableI)
-      fix a assume "a \<in> {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
-      then obtain A where "a = X i -` A \<inter> space M" "A \<in> sets (M' i)" by auto
+      fix a assume "a \<in> {X i -` A \<inter> space M |A. A \<in> E i}"
+      then obtain A where "a = X i -` A \<inter> space M" "A \<in> E i" by auto
       moreover
-      fix b assume "b \<in> {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
-      then obtain B where "b = X i -` B \<inter> space M" "B \<in> sets (M' i)" by auto
+      fix b assume "b \<in> {X i -` A \<inter> space M |A. A \<in> E i}"
+      then obtain B where "b = X i -` B \<inter> space M" "B \<in> E i" by auto
       moreover
       have "(X i -` A \<inter> space M) \<inter> (X i -` B \<inter> space M) = X i -` (A \<inter> B) \<inter> space M" by auto
       moreover note Int_stable[OF `i \<in> I`]
       ultimately
-      show "a \<inter> b \<in> {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
+      show "a \<inter> b \<in> {X i -` A \<inter> space M |A. A \<in> E i}"
         by (auto simp del: vimage_Int intro!: exI[of _ "A \<inter> B"] dest: Int_stableD)
     qed }
-  note indep_sets_sigma_sets_iff[OF this, simp]
+  note indep_sets_X = indep_sets_sigma_sets_iff[OF this]
 
   { fix i assume "i \<in> I"
-    { fix A assume "A \<in> sets (M' i)"
-      then have "A \<in> sets (sigma (M' i))" by (auto simp: sets_sigma intro: sigma_sets.Basic)
+    { fix A assume "A \<in> E i"
+      with M'[OF `i \<in> I`] have "A \<in> sets (M' i)" by auto
       moreover
-      from rv[OF `i\<in>I`] have "X i \<in> measurable M (sigma (M' i))" by auto
+      from rv[OF `i\<in>I`] have "X i \<in> measurable M (M' i)" by auto
       ultimately
       have "X i -` A \<inter> space M \<in> sets M" by (auto intro: measurable_sets) }
     with X[OF `i\<in>I`] space[OF `i\<in>I`]
-    have "{X i -` A \<inter> space M |A. A \<in> sets (M' i)} \<subseteq> events"
-      "space M \<in> {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
+    have "{X i -` A \<inter> space M |A. A \<in> E i} \<subseteq> events"
+      "space M \<in> {X i -` A \<inter> space M |A. A \<in> E i}"
       by (auto intro!: exI[of _ "space (M' i)"]) }
-  note indep_sets_finite[OF I this, simp]
+  note indep_sets_finite_X = indep_sets_finite[OF I this]
 
-  have "(\<forall>A\<in>\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> sets (M' i)}. prob (INTER I A) = (\<Prod>j\<in>I. prob (A j))) =
-    (\<forall>A\<in>\<Pi> i\<in>I. sets (M' i). prob ((\<Inter>j\<in>I. X j -` A j) \<inter> space M) = (\<Prod>x\<in>I. prob (X x -` A x \<inter> space M)))"
+  have "(\<forall>A\<in>\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> E i}. prob (INTER I A) = (\<Prod>j\<in>I. prob (A j))) =
+    (\<forall>A\<in>\<Pi> i\<in>I. E i. prob ((\<Inter>j\<in>I. X j -` A j) \<inter> space M) = (\<Prod>x\<in>I. prob (X x -` A x \<inter> space M)))"
     (is "?L = ?R")
   proof safe
-    fix A assume ?L and A: "A \<in> (\<Pi> i\<in>I. sets (M' i))"
+    fix A assume ?L and A: "A \<in> (\<Pi> i\<in>I. E i)"
     from `?L`[THEN bspec, of "\<lambda>i. X i -` A i \<inter> space M"] A `I \<noteq> {}`
     show "prob ((\<Inter>j\<in>I. X j -` A j) \<inter> space M) = (\<Prod>x\<in>I. prob (X x -` A x \<inter> space M))"
       by (auto simp add: Pi_iff)
   next
-    fix A assume ?R and A: "A \<in> (\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> sets (M' i)})"
-    from A have "\<forall>i\<in>I. \<exists>B. A i = X i -` B \<inter> space M \<and> B \<in> sets (M' i)" by auto
+    fix A assume ?R and A: "A \<in> (\<Pi> i\<in>I. {X i -` A \<inter> space M |A. A \<in> E i})"
+    from A have "\<forall>i\<in>I. \<exists>B. A i = X i -` B \<inter> space M \<and> B \<in> E i" by auto
     from bchoice[OF this] obtain B where B: "\<forall>i\<in>I. A i = X i -` B i \<inter> space M"
-      "B \<in> (\<Pi> i\<in>I. sets (M' i))" by auto
+      "B \<in> (\<Pi> i\<in>I. E i)" by auto
     from `?R`[THEN bspec, OF B(2)] B(1) `I \<noteq> {}`
     show "prob (INTER I A) = (\<Prod>j\<in>I. prob (A j))"
       by simp
   qed
   then show ?thesis using `I \<noteq> {}`
-    by (simp add: rv indep_vars_def)
+    by (simp add: rv indep_vars_def indep_sets_X sigma_sets_X indep_sets_finite_X cong: indep_sets_cong)
 qed
 
 lemma (in prob_space) indep_vars_compose:
   assumes "indep_vars M' X I"
-  assumes rv:
-    "\<And>i. i \<in> I \<Longrightarrow> sigma_algebra (N i)"
-    "\<And>i. i \<in> I \<Longrightarrow> Y i \<in> measurable (M' i) (N i)"
+  assumes rv: "\<And>i. i \<in> I \<Longrightarrow> Y i \<in> measurable (M' i) (N i)"
   shows "indep_vars N (\<lambda>i. Y i \<circ> X i) I"
   unfolding indep_vars_def
 proof
   from rv `indep_vars M' X I`
   show "\<forall>i\<in>I. random_variable (N i) (Y i \<circ> X i)"
-    by (auto intro!: measurable_comp simp: indep_vars_def)
+    by (auto simp: indep_vars_def)
 
   have "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I"
     using `indep_vars M' X I` by (simp add: indep_vars_def)
@@ -806,7 +792,7 @@
   qed
 qed
 
-lemma (in prob_space) indep_varsD:
+lemma (in prob_space) indep_varsD_finite:
   assumes X: "indep_vars M' X I"
   assumes I: "I \<noteq> {}" "finite I" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M' i)"
   shows "prob (\<Inter>i\<in>I. X i -` A i \<inter> space M) = (\<Prod>i\<in>I. prob (X i -` A i \<inter> space M))"
@@ -815,96 +801,134 @@
     using X by (auto simp: indep_vars_def)
   show "I \<subseteq> I" "I \<noteq> {}" "finite I" using I by auto
   show "\<forall>i\<in>I. X i -` A i \<inter> space M \<in> sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
-    using I by (auto intro: sigma_sets.Basic)
+    using I by auto
 qed
 
-lemma (in prob_space) indep_distribution_eq_measure:
-  assumes I: "I \<noteq> {}" "finite I"
+lemma (in prob_space) indep_varsD:
+  assumes X: "indep_vars M' X I"
+  assumes I: "J \<noteq> {}" "finite J" "J \<subseteq> I" "\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M' i)"
+  shows "prob (\<Inter>i\<in>J. X i -` A i \<inter> space M) = (\<Prod>i\<in>J. prob (X i -` A i \<inter> space M))"
+proof (rule indep_setsD)
+  show "indep_sets (\<lambda>i. sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}) I"
+    using X by (auto simp: indep_vars_def)
+  show "\<forall>i\<in>J. X i -` A i \<inter> space M \<in> sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)}"
+    using I by auto
+qed fact+
+
+lemma prod_algebra_cong:
+  assumes "I = J" and sets: "(\<And>i. i \<in> I \<Longrightarrow> sets (M i) = sets (N i))"
+  shows "prod_algebra I M = prod_algebra J N"
+proof -
+  have space: "\<And>i. i \<in> I \<Longrightarrow> space (M i) = space (N i)"
+    using sets_eq_imp_space_eq[OF sets] by auto
+  with sets show ?thesis unfolding `I = J`
+    by (intro antisym prod_algebra_mono) auto
+qed
+
+lemma space_in_prod_algebra:
+  "(\<Pi>\<^isub>E i\<in>I. space (M i)) \<in> prod_algebra I M"
+proof cases
+  assume "I = {}" then show ?thesis
+    by (auto simp add: prod_algebra_def image_iff prod_emb_def)
+next
+  assume "I \<noteq> {}"
+  then obtain i where "i \<in> I" by auto
+  then have "(\<Pi>\<^isub>E i\<in>I. space (M i)) = prod_emb I M {i} (\<Pi>\<^isub>E i\<in>{i}. space (M i))"
+    by (auto simp: prod_emb_def Pi_iff)
+  also have "\<dots> \<in> prod_algebra I M"
+    using `i \<in> I` by (intro prod_algebraI) auto
+  finally show ?thesis .
+qed
+
+lemma (in prob_space) indep_vars_iff_distr_eq_PiM:
+  fixes I :: "'i set" and X :: "'i \<Rightarrow> 'a \<Rightarrow> 'b"
+  assumes "I \<noteq> {}"
   assumes rv: "\<And>i. random_variable (M' i) (X i)"
   shows "indep_vars M' X I \<longleftrightarrow>
-    (\<forall>A\<in>sets (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := ereal\<circ>distribution (X i) \<rparr>)).
-      distribution (\<lambda>x. \<lambda>i\<in>I. X i x) A =
-      finite_measure.\<mu>' (\<Pi>\<^isub>M i\<in>I. (M' i \<lparr> measure := ereal\<circ>distribution (X i) \<rparr>)) A)"
-    (is "_ \<longleftrightarrow> (\<forall>X\<in>_. distribution ?D X = finite_measure.\<mu>' (Pi\<^isub>M I ?M) X)")
+    distr M (\<Pi>\<^isub>M i\<in>I. M' i) (\<lambda>x. \<lambda>i\<in>I. X i x) = (\<Pi>\<^isub>M i\<in>I. distr M (M' i) (X i))"
 proof -
-  interpret M': prob_space "?M i" for i
-    using rv by (rule distribution_prob_space)
-  interpret P: finite_product_prob_space ?M I
-    proof qed fact
+  let ?P = "\<Pi>\<^isub>M i\<in>I. M' i"
+  let ?X = "\<lambda>x. \<lambda>i\<in>I. X i x"
+  let ?D = "distr M ?P ?X"
+  have X: "random_variable ?P ?X" by (intro measurable_restrict rv)
+  interpret D: prob_space ?D by (intro prob_space_distr X)
 
-  let ?D' = "(Pi\<^isub>M I ?M) \<lparr> measure := ereal \<circ> distribution ?D \<rparr>"
-  have "random_variable P.P ?D"
-    using `finite I` rv by (intro random_variable_restrict) auto
-  then interpret D: prob_space ?D'
-    by (rule distribution_prob_space)
-
+  let ?D' = "\<lambda>i. distr M (M' i) (X i)"
+  let ?P' = "\<Pi>\<^isub>M i\<in>I. distr M (M' i) (X i)"
+  interpret D': prob_space "?D' i" for i by (intro prob_space_distr rv)
+  interpret P: product_prob_space ?D' I ..
+    
   show ?thesis
-  proof (intro iffI ballI)
+  proof
     assume "indep_vars M' X I"
-    fix A assume "A \<in> sets P.P"
-    moreover
-    have "D.prob A = P.prob A"
-    proof (rule prob_space_unique_Int_stable)
-      show "prob_space ?D'" by unfold_locales
-      show "prob_space (Pi\<^isub>M I ?M)" by unfold_locales
-      show "Int_stable P.G" using M'.Int
-        by (intro Int_stable_product_algebra_generator) (simp add: Int_stable_def)
-      show "space P.G \<in> sets P.G"
-        using M'.top by (simp add: product_algebra_generator_def)
-      show "space ?D' = space P.G"  "sets ?D' = sets (sigma P.G)"
-        by (simp_all add: product_algebra_def product_algebra_generator_def sets_sigma)
-      show "space P.P = space P.G" "sets P.P = sets (sigma P.G)"
-        by (simp_all add: product_algebra_def)
-      show "A \<in> sets (sigma P.G)"
-        using `A \<in> sets P.P` by (simp add: product_algebra_def)
+    show "?D = ?P'"
+    proof (rule measure_eqI_generator_eq)
+      show "Int_stable (prod_algebra I M')"
+        by (rule Int_stable_prod_algebra)
+      show "prod_algebra I M' \<subseteq> Pow (space ?P)"
+        using prod_algebra_sets_into_space by (simp add: space_PiM)
+      show "sets ?D = sigma_sets (space ?P) (prod_algebra I M')"
+        by (simp add: sets_PiM space_PiM)
+      show "sets ?P' = sigma_sets (space ?P) (prod_algebra I M')"
+        by (simp add: sets_PiM space_PiM cong: prod_algebra_cong)
+      let ?A = "\<lambda>i. \<Pi>\<^isub>E i\<in>I. space (M' i)"
+      show "range ?A \<subseteq> prod_algebra I M'" "incseq ?A" "(\<Union>i. ?A i) = space (Pi\<^isub>M I M')"
+        by (auto simp: space_PiM intro!: space_in_prod_algebra cong: prod_algebra_cong)
+      { fix i show "emeasure ?D (\<Pi>\<^isub>E i\<in>I. space (M' i)) \<noteq> \<infinity>" by auto }
+    next
+      fix E assume E: "E \<in> prod_algebra I M'"
+      from prod_algebraE[OF E] guess J Y . note J = this
 
-      fix E assume E: "E \<in> sets P.G"
-      then have "E \<in> sets P.P"
-        by (simp add: sets_sigma sigma_sets.Basic product_algebra_def)
-      then have "D.prob E = distribution ?D E"
-        unfolding D.\<mu>'_def by simp
-      also
-      from E obtain F where "E = Pi\<^isub>E I F" and F: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets (M' i)"
-        by (auto simp: product_algebra_generator_def)
-      with `I \<noteq> {}` have "distribution ?D E = prob (\<Inter>i\<in>I. X i -` F i \<inter> space M)"
-        using `I \<noteq> {}` by (auto intro!: arg_cong[where f=prob] simp: Pi_iff distribution_def)
-      also have "\<dots> = (\<Prod>i\<in>I. prob (X i -` F i \<inter> space M))"
-        using `indep_vars M' X I` I F by (rule indep_varsD)
-      also have "\<dots> = P.prob E"
-        using F by (simp add: `E = Pi\<^isub>E I F` P.prob_times M'.\<mu>'_def distribution_def)
-      finally show "D.prob E = P.prob E" .
+      from E have "E \<in> sets ?P" by (auto simp: sets_PiM)
+      then have "emeasure ?D E = emeasure M (?X -` E \<inter> space M)"
+        by (simp add: emeasure_distr X)
+      also have "?X -` E \<inter> space M = (\<Inter>i\<in>J. X i -` Y i \<inter> space M)"
+        using J `I \<noteq> {}` measurable_space[OF rv] by (auto simp: prod_emb_def Pi_iff split: split_if_asm)
+      also have "emeasure M (\<Inter>i\<in>J. X i -` Y i \<inter> space M) = (\<Prod> i\<in>J. emeasure M (X i -` Y i \<inter> space M))"
+        using `indep_vars M' X I` J `I \<noteq> {}` using indep_varsD[of M' X I J]
+        by (auto simp: emeasure_eq_measure setprod_ereal)
+      also have "\<dots> = (\<Prod> i\<in>J. emeasure (?D' i) (Y i))"
+        using rv J by (simp add: emeasure_distr)
+      also have "\<dots> = emeasure ?P' E"
+        using P.emeasure_PiM_emb[of J Y] J by (simp add: prod_emb_def)
+      finally show "emeasure ?D E = emeasure ?P' E" .
     qed
-    ultimately show "distribution ?D A = P.prob A"
-      by (simp add: D.\<mu>'_def)
   next
-    assume eq: "\<forall>A\<in>sets P.P. distribution ?D A = P.prob A"
-    have [simp]: "\<And>i. sigma (M' i) = M' i"
-      using rv by (intro sigma_algebra.sigma_eq) simp
-    have "indep_vars (\<lambda>i. sigma (M' i)) X I"
-    proof (subst indep_vars_finite[OF I])
-      fix i assume [simp]: "i \<in> I"
-      show "random_variable (sigma (M' i)) (X i)"
-        using rv[of i] by simp
-      show "Int_stable (M' i)" "space (M' i) \<in> sets (M' i)"
-        using M'.Int[of _ i] M'.top by (auto simp: Int_stable_def)
+    assume "?D = ?P'"
+    show "indep_vars M' X I" unfolding indep_vars_def
+    proof (intro conjI indep_setsI ballI rv)
+      fix i show "sigma_sets (space M) {X i -` A \<inter> space M |A. A \<in> sets (M' i)} \<subseteq> events"
+        by (auto intro!: sigma_sets_subset measurable_sets rv)
     next
-      show "\<forall>A\<in>\<Pi> i\<in>I. sets (M' i). prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = (\<Prod>j\<in>I. prob (X j -` A j \<inter> space M))"
+      fix J Y' assume J: "J \<noteq> {}" "J \<subseteq> I" "finite J"
+      assume Y': "\<forall>j\<in>J. Y' j \<in> sigma_sets (space M) {X j -` A \<inter> space M |A. A \<in> sets (M' j)}"
+      have "\<forall>j\<in>J. \<exists>Y. Y' j = X j -` Y \<inter> space M \<and> Y \<in> sets (M' j)"
       proof
-        fix A assume A: "A \<in> (\<Pi> i\<in>I. sets (M' i))"
-        then have A_in_P: "(Pi\<^isub>E I A) \<in> sets P.P"
-          by (auto intro!: product_algebraI)
-        have "prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = distribution ?D (Pi\<^isub>E I A)"
-          using `I \<noteq> {}`by (auto intro!: arg_cong[where f=prob] simp: Pi_iff distribution_def)
-        also have "\<dots> = P.prob (Pi\<^isub>E I A)" using A_in_P eq by simp
-        also have "\<dots> = (\<Prod>i\<in>I. M'.prob i (A i))"
-          using A by (intro P.prob_times) auto
-        also have "\<dots> = (\<Prod>i\<in>I. prob (X i -` A i \<inter> space M))"
-          using A by (auto intro!: setprod_cong simp: M'.\<mu>'_def Pi_iff distribution_def)
-        finally show "prob (\<Inter>j\<in>I. X j -` A j \<inter> space M) = (\<Prod>j\<in>I. prob (X j -` A j \<inter> space M))" .
+        fix j assume "j \<in> J"
+        from Y'[rule_format, OF this] rv[of j]
+        show "\<exists>Y. Y' j = X j -` Y \<inter> space M \<and> Y \<in> sets (M' j)"
+          by (subst (asm) sigma_sets_vimage_commute[symmetric, of _ _ "space (M' j)"])
+             (auto dest: measurable_space simp: sigma_sets_eq)
       qed
+      from bchoice[OF this] obtain Y where
+        Y: "\<And>j. j \<in> J \<Longrightarrow> Y' j = X j -` Y j \<inter> space M" "\<And>j. j \<in> J \<Longrightarrow> Y j \<in> sets (M' j)" by auto
+      let ?E = "prod_emb I M' J (Pi\<^isub>E J Y)"
+      from Y have "(\<Inter>j\<in>J. Y' j) = ?X -` ?E \<inter> space M"
+        using J `I \<noteq> {}` measurable_space[OF rv] by (auto simp: prod_emb_def Pi_iff split: split_if_asm)
+      then have "emeasure M (\<Inter>j\<in>J. Y' j) = emeasure M (?X -` ?E \<inter> space M)"
+        by simp
+      also have "\<dots> = emeasure ?D ?E"
+        using Y  J by (intro emeasure_distr[symmetric] X sets_PiM_I) auto
+      also have "\<dots> = emeasure ?P' ?E"
+        using `?D = ?P'` by simp
+      also have "\<dots> = (\<Prod> i\<in>J. emeasure (?D' i) (Y i))"
+        using P.emeasure_PiM_emb[of J Y] J Y by (simp add: prod_emb_def)
+      also have "\<dots> = (\<Prod> i\<in>J. emeasure M (Y' i))"
+        using rv J Y by (simp add: emeasure_distr)
+      finally have "emeasure M (\<Inter>j\<in>J. Y' j) = (\<Prod> i\<in>J. emeasure M (Y' i))" .
+      then show "prob (\<Inter>j\<in>J. Y' j) = (\<Prod> i\<in>J. prob (Y' i))"
+        by (auto simp: emeasure_eq_measure setprod_ereal)
     qed
-    then show "indep_vars M' X I"
-      by simp
   qed
 qed
 
@@ -936,56 +960,188 @@
     unfolding UNIV_bool by auto
 qed
 
-lemma (in prob_space) indep_var_distributionD:
-  assumes indep: "indep_var S X T Y"
-  defines "P \<equiv> S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
-  assumes "A \<in> sets P"
-  shows "joint_distribution X Y A = finite_measure.\<mu>' P A"
-proof -
-  from indep have rvs: "random_variable S X" "random_variable T Y"
-    by (blast dest: indep_var_rv1 indep_var_rv2)+
+lemma measurable_bool_case[simp, intro]:
+  "(\<lambda>(x, y). bool_case x y) \<in> measurable (M \<Otimes>\<^isub>M N) (Pi\<^isub>M UNIV (bool_case M N))"
+    (is "?f \<in> measurable ?B ?P")
+proof (rule measurable_PiM_single)
+  show "?f \<in> space ?B \<rightarrow> (\<Pi>\<^isub>E i\<in>UNIV. space (bool_case M N i))"
+    by (auto simp: space_pair_measure extensional_def split: bool.split)
+  fix i A assume "A \<in> sets (case i of True \<Rightarrow> M | False \<Rightarrow> N)"
+  moreover then have "{\<omega> \<in> space (M \<Otimes>\<^isub>M N). prod_case bool_case \<omega> i \<in> A}
+    = (case i of True \<Rightarrow> A \<times> space N | False \<Rightarrow> space M \<times> A)" 
+    by (auto simp: space_pair_measure split: bool.split dest!: sets_into_space)
+  ultimately show "{\<omega> \<in> space (M \<Otimes>\<^isub>M N). prod_case bool_case \<omega> i \<in> A} \<in> sets ?B"
+    by (auto split: bool.split)
+qed
+
+lemma borel_measurable_indicator':
+  "A \<in> sets N \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> (\<lambda>x. indicator A (f x)) \<in> borel_measurable M"
+  using measurable_comp[OF _ borel_measurable_indicator, of f M N A] by (auto simp add: comp_def)
+
+lemma (in product_sigma_finite) distr_component:
+  "distr (M i) (Pi\<^isub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^isub>M {i} M" (is "?D = ?P")
+proof (intro measure_eqI[symmetric])
+  interpret I: finite_product_sigma_finite M "{i}" by default simp
+
+  have eq: "\<And>x. x \<in> extensional {i} \<Longrightarrow> (\<lambda>j\<in>{i}. x i) = x"
+    by (auto simp: extensional_def restrict_def)
+
+  fix A assume A: "A \<in> sets ?P"
+  then have "emeasure ?P A = (\<integral>\<^isup>+x. indicator A x \<partial>?P)" 
+    by simp
+  also have "\<dots> = (\<integral>\<^isup>+x. indicator ((\<lambda>x. \<lambda>i\<in>{i}. x) -` A \<inter> space (M i)) x \<partial>M i)" 
+    apply (subst product_positive_integral_singleton[symmetric])
+    apply (force intro!: measurable_restrict measurable_sets A)
+    apply (auto intro!: positive_integral_cong simp: space_PiM indicator_def simp: eq)
+    done
+  also have "\<dots> = emeasure (M i) ((\<lambda>x. \<lambda>i\<in>{i}. x) -` A \<inter> space (M i))"
+    by (force intro!: measurable_restrict measurable_sets A positive_integral_indicator)
+  also have "\<dots> = emeasure ?D A"
+    using A by (auto intro!: emeasure_distr[symmetric] measurable_restrict) 
+  finally show "emeasure (Pi\<^isub>M {i} M) A = emeasure ?D A" .
+qed simp
 
-  let ?S = "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
-  let ?T = "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
-  interpret X: prob_space ?S by (rule distribution_prob_space) fact
-  interpret Y: prob_space ?T by (rule distribution_prob_space) fact
-  interpret XY: pair_prob_space ?S ?T by default
+lemma pair_measure_eqI:
+  assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
+  assumes sets: "sets (M1 \<Otimes>\<^isub>M M2) = sets M"
+  assumes emeasure: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> emeasure M1 A * emeasure M2 B = emeasure M (A \<times> B)"
+  shows "M1 \<Otimes>\<^isub>M M2 = M"
+proof -
+  interpret M1: sigma_finite_measure M1 by fact
+  interpret M2: sigma_finite_measure M2 by fact
+  interpret pair_sigma_finite M1 M2 by default
+  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
+  let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
+  let ?P = "M1 \<Otimes>\<^isub>M M2"
+  show ?thesis
+  proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])
+    show "?E \<subseteq> Pow (space ?P)"
+      using space_closed[of M1] space_closed[of M2] by (auto simp: space_pair_measure)
+    show "sets ?P = sigma_sets (space ?P) ?E"
+      by (simp add: sets_pair_measure space_pair_measure)
+    then show "sets M = sigma_sets (space ?P) ?E"
+      using sets[symmetric] by simp
+  next
+    show "range F \<subseteq> ?E" "incseq F" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"
+      using F by (auto simp: space_pair_measure)
+  next
+    fix X assume "X \<in> ?E"
+    then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto
+    then have "emeasure ?P X = emeasure M1 A * emeasure M2 B"
+       by (simp add: emeasure_pair_measure_Times)
+    also have "\<dots> = emeasure M (A \<times> B)"
+      using A B emeasure by auto
+    finally show "emeasure ?P X = emeasure M X"
+      by simp
+  qed
+qed
 
-  let ?J = "XY.P\<lparr> measure := ereal \<circ> joint_distribution X Y \<rparr>"
-  interpret J: prob_space ?J
-    by (rule joint_distribution_prob_space) (simp_all add: rvs)
+lemma pair_measure_eq_distr_PiM:
+  fixes M1 :: "'a measure" and M2 :: "'a measure"
+  assumes "sigma_finite_measure M1" "sigma_finite_measure M2"
+  shows "(M1 \<Otimes>\<^isub>M M2) = distr (Pi\<^isub>M UNIV (bool_case M1 M2)) (M1 \<Otimes>\<^isub>M M2) (\<lambda>x. (x True, x False))"
+    (is "?P = ?D")
+proof (rule pair_measure_eqI[OF assms])
+  interpret B: product_sigma_finite "bool_case M1 M2"
+    unfolding product_sigma_finite_def using assms by (auto split: bool.split)
+  let ?B = "Pi\<^isub>M UNIV (bool_case M1 M2)"
 
-  have "finite_measure.\<mu>' (XY.P\<lparr> measure := ereal \<circ> joint_distribution X Y \<rparr>) A = XY.\<mu>' A"
-  proof (rule prob_space_unique_Int_stable)
-    show "Int_stable (pair_measure_generator ?S ?T)" (is "Int_stable ?P")
-      by fact
-    show "space ?P \<in> sets ?P"
-      unfolding space_pair_measure[simplified pair_measure_def space_sigma]
-      using X.top Y.top by (auto intro!: pair_measure_generatorI)
+  have [simp]: "fst \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x True)" "snd \<circ> (\<lambda>x. (x True, x False)) = (\<lambda>x. x False)"
+    by auto
+  fix A B assume A: "A \<in> sets M1" and B: "B \<in> sets M2"
+  have "emeasure M1 A * emeasure M2 B = (\<Prod> i\<in>UNIV. emeasure (bool_case M1 M2 i) (bool_case A B i))"
+    by (simp add: UNIV_bool ac_simps)
+  also have "\<dots> = emeasure ?B (Pi\<^isub>E UNIV (bool_case A B))"
+    using A B by (subst B.emeasure_PiM) (auto split: bool.split)
+  also have "Pi\<^isub>E UNIV (bool_case A B) = (\<lambda>x. (x True, x False)) -` (A \<times> B) \<inter> space ?B"
+    using A[THEN sets_into_space] B[THEN sets_into_space]
+    by (auto simp: Pi_iff all_bool_eq space_PiM split: bool.split)
+  finally show "emeasure M1 A * emeasure M2 B = emeasure ?D (A \<times> B)"
+    using A B
+      measurable_component_singleton[of True UNIV "bool_case M1 M2"]
+      measurable_component_singleton[of False UNIV "bool_case M1 M2"]
+    by (subst emeasure_distr) (auto simp: measurable_pair_iff)
+qed simp
 
-    show "prob_space ?J" by unfold_locales
-    show "space ?J = space ?P"
-      by (simp add: pair_measure_generator_def space_pair_measure)
-    show "sets ?J = sets (sigma ?P)"
-      by (simp add: pair_measure_def)
+lemma measurable_Pair:
+  assumes rvs: "X \<in> measurable M S" "Y \<in> measurable M T"
+  shows "(\<lambda>x. (X x, Y x)) \<in> measurable M (S \<Otimes>\<^isub>M T)"
+proof -
+  have [simp]: "fst \<circ> (\<lambda>x. (X x, Y x)) = (\<lambda>x. X x)" "snd \<circ> (\<lambda>x. (X x, Y x)) = (\<lambda>x. Y x)"
+    by auto
+  show " (\<lambda>x. (X x, Y x)) \<in> measurable M (S \<Otimes>\<^isub>M T)"
+    by (auto simp: measurable_pair_iff rvs)
+qed
+
+lemma (in prob_space) indep_var_distribution_eq:
+  "indep_var S X T Y \<longleftrightarrow> random_variable S X \<and> random_variable T Y \<and>
+    distr M S X \<Otimes>\<^isub>M distr M T Y = distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))" (is "_ \<longleftrightarrow> _ \<and> _ \<and> ?S \<Otimes>\<^isub>M ?T = ?J")
+proof safe
+  assume "indep_var S X T Y"
+  then show rvs: "random_variable S X" "random_variable T Y"
+    by (blast dest: indep_var_rv1 indep_var_rv2)+
+  then have XY: "random_variable (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+    by (rule measurable_Pair)
+
+  interpret X: prob_space ?S by (rule prob_space_distr) fact
+  interpret Y: prob_space ?T by (rule prob_space_distr) fact
+  interpret XY: pair_prob_space ?S ?T ..
+  show "?S \<Otimes>\<^isub>M ?T = ?J"
+  proof (rule pair_measure_eqI)
+    show "sigma_finite_measure ?S" ..
+    show "sigma_finite_measure ?T" ..
 
-    show "prob_space XY.P" by unfold_locales
-    show "space XY.P = space ?P" "sets XY.P = sets (sigma ?P)"
-      by (simp_all add: pair_measure_generator_def pair_measure_def)
+    fix A B assume A: "A \<in> sets ?S" and B: "B \<in> sets ?T"
+    have "emeasure ?J (A \<times> B) = emeasure M ((\<lambda>x. (X x, Y x)) -` (A \<times> B) \<inter> space M)"
+      using A B by (intro emeasure_distr[OF XY]) auto
+    also have "\<dots> = emeasure M (X -` A \<inter> space M) * emeasure M (Y -` B \<inter> space M)"
+      using indep_varD[OF `indep_var S X T Y`, of A B] A B by (simp add: emeasure_eq_measure)
+    also have "\<dots> = emeasure ?S A * emeasure ?T B"
+      using rvs A B by (simp add: emeasure_distr)
+    finally show "emeasure ?S A * emeasure ?T B = emeasure ?J (A \<times> B)" by simp
+  qed simp
+next
+  assume rvs: "random_variable S X" "random_variable T Y"
+  then have XY: "random_variable (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+    by (rule measurable_Pair)
 
-    show "A \<in> sets (sigma ?P)"
-      using `A \<in> sets P` unfolding P_def pair_measure_def by simp
+  let ?S = "distr M S X" and ?T = "distr M T Y"
+  interpret X: prob_space ?S by (rule prob_space_distr) fact
+  interpret Y: prob_space ?T by (rule prob_space_distr) fact
+  interpret XY: pair_prob_space ?S ?T ..
+
+  assume "?S \<Otimes>\<^isub>M ?T = ?J"
 
-    fix X assume "X \<in> sets ?P"
-    then obtain A B where "A \<in> sets S" "B \<in> sets T" "X = A \<times> B"
-      by (auto simp: sets_pair_measure_generator)
-    then show "J.\<mu>' X = XY.\<mu>' X"
-      unfolding J.\<mu>'_def XY.\<mu>'_def using indep
-      by (simp add: XY.pair_measure_times)
-         (simp add: distribution_def indep_varD)
+  { fix S and X
+    have "Int_stable {X -` A \<inter> space M |A. A \<in> sets S}"
+    proof (safe intro!: Int_stableI)
+      fix A B assume "A \<in> sets S" "B \<in> sets S"
+      then show "\<exists>C. (X -` A \<inter> space M) \<inter> (X -` B \<inter> space M) = (X -` C \<inter> space M) \<and> C \<in> sets S"
+        by (intro exI[of _ "A \<inter> B"]) auto
+    qed }
+  note Int_stable = this
+
+  show "indep_var S X T Y" unfolding indep_var_eq
+  proof (intro conjI indep_set_sigma_sets Int_stable rvs)
+    show "indep_set {X -` A \<inter> space M |A. A \<in> sets S} {Y -` A \<inter> space M |A. A \<in> sets T}"
+    proof (safe intro!: indep_setI)
+      { fix A assume "A \<in> sets S" then show "X -` A \<inter> space M \<in> sets M"
+        using `X \<in> measurable M S` by (auto intro: measurable_sets) }
+      { fix A assume "A \<in> sets T" then show "Y -` A \<inter> space M \<in> sets M"
+        using `Y \<in> measurable M T` by (auto intro: measurable_sets) }
+    next
+      fix A B assume ab: "A \<in> sets S" "B \<in> sets T"
+      then have "ereal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) = emeasure ?J (A \<times> B)"
+        using XY by (auto simp add: emeasure_distr emeasure_eq_measure intro!: arg_cong[where f="prob"])
+      also have "\<dots> = emeasure (?S \<Otimes>\<^isub>M ?T) (A \<times> B)"
+        unfolding `?S \<Otimes>\<^isub>M ?T = ?J` ..
+      also have "\<dots> = emeasure ?S A * emeasure ?T B"
+        using ab by (simp add: XY.emeasure_pair_measure_Times)
+      finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) =
+        prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)"
+        using rvs ab by (simp add: emeasure_eq_measure emeasure_distr)
+    qed
   qed
-  then show ?thesis
-    using `A \<in> sets P` unfolding P_def J.\<mu>'_def XY.\<mu>'_def by simp
 qed
 
 end
--- a/src/HOL/Probability/Infinite_Product_Measure.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -5,9 +5,49 @@
 header {*Infinite Product Measure*}
 
 theory Infinite_Product_Measure
-  imports Probability_Measure
+  imports Probability_Measure Caratheodory
 begin
 
+lemma sigma_sets_mono: assumes "A \<subseteq> sigma_sets X B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
+proof
+  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
+    by induct (insert `A \<subseteq> sigma_sets X B`, auto intro: sigma_sets.intros)
+qed
+
+lemma sigma_sets_mono': assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
+proof
+  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
+    by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros)
+qed
+
+lemma sigma_sets_superset_generator: "A \<subseteq> sigma_sets X A"
+  by (auto intro: sigma_sets.Basic)
+
+lemma (in product_sigma_finite)
+  assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
+  shows emeasure_fold_integral:
+    "emeasure (Pi\<^isub>M (I \<union> J) M) A = (\<integral>\<^isup>+x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M)) \<partial>Pi\<^isub>M I M)" (is ?I)
+    and emeasure_fold_measurable:
+    "(\<lambda>x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M))) \<in> borel_measurable (Pi\<^isub>M I M)" (is ?B)
+proof -
+  interpret I: finite_product_sigma_finite M I by default fact
+  interpret J: finite_product_sigma_finite M J by default fact
+  interpret IJ: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" ..
+  have merge: "(\<lambda>(x, y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
+    by (intro measurable_sets[OF _ A] measurable_merge assms)
+
+  show ?I
+    apply (subst distr_merge[symmetric, OF IJ])
+    apply (subst emeasure_distr[OF measurable_merge[OF IJ(1)] A])
+    apply (subst IJ.emeasure_pair_measure_alt[OF merge])
+    apply (auto intro!: positive_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure)
+    done
+
+  show ?B
+    using IJ.measurable_emeasure_Pair1[OF merge]
+    by (simp add: vimage_compose[symmetric] comp_def space_pair_measure cong: measurable_cong)
+qed
+
 lemma restrict_extensional_sub[intro]: "A \<subseteq> B \<Longrightarrow> restrict f A \<in> extensional B"
   unfolding restrict_def extensional_def by auto
 
@@ -41,189 +81,178 @@
   qed
 qed
 
-lemma (in product_prob_space) measure_preserving_restrict:
+lemma prod_algebraI_finite:
+  "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^isub>E I E) \<in> prod_algebra I M"
+  using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets_into_space] by simp
+
+lemma Int_stable_PiE: "Int_stable {Pi\<^isub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
+proof (safe intro!: Int_stableI)
+  fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
+  then show "\<exists>G. Pi\<^isub>E J E \<inter> Pi\<^isub>E J F = Pi\<^isub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
+    by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"])
+qed
+
+lemma prod_emb_trans[simp]:
+  "J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X"
+  by (auto simp add: Int_absorb1 prod_emb_def)
+
+lemma prod_emb_Pi:
+  assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
+  shows "prod_emb K M J (Pi\<^isub>E J X) = (\<Pi>\<^isub>E i\<in>K. if i \<in> J then X i else space (M i))"
+  using assms space_closed
+  by (auto simp: prod_emb_def Pi_iff split: split_if_asm) blast+
+
+lemma prod_emb_id:
+  "B \<subseteq> (\<Pi>\<^isub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
+  by (auto simp: prod_emb_def Pi_iff subset_eq extensional_restrict)
+
+lemma measurable_prod_emb[intro, simp]:
+  "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^isub>M L M)"
+  unfolding prod_emb_def space_PiM[symmetric]
+  by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton)
+
+lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^isub>M L M) (Pi\<^isub>M J M)"
+  by (intro measurable_restrict measurable_component_singleton) auto
+
+lemma (in product_prob_space) distr_restrict:
   assumes "J \<noteq> {}" "J \<subseteq> K" "finite K"
-  shows "(\<lambda>f. restrict f J) \<in> measure_preserving (\<Pi>\<^isub>M i\<in>K. M i) (\<Pi>\<^isub>M i\<in>J. M i)" (is "?R \<in> _")
-proof -
-  interpret K: finite_product_prob_space M K by default fact
-  have J: "J \<noteq> {}" "finite J" using assms by (auto simp add: finite_subset)
-  interpret J: finite_product_prob_space M J
-    by default (insert J, auto)
-  from J.sigma_finite_pairs guess F .. note F = this
-  then have [simp,intro]: "\<And>k i. k \<in> J \<Longrightarrow> F k i \<in> sets (M k)"
-    by auto
-  let ?F = "\<lambda>i. \<Pi>\<^isub>E k\<in>J. F k i"
-  let ?J = "product_algebra_generator J M \<lparr> measure := measure (Pi\<^isub>M J M) \<rparr>"
-  have "?R \<in> measure_preserving (\<Pi>\<^isub>M i\<in>K. M i) (sigma ?J)"
-  proof (rule K.measure_preserving_Int_stable)
-    show "Int_stable ?J"
-      by (auto simp: Int_stable_def product_algebra_generator_def PiE_Int)
-    show "range ?F \<subseteq> sets ?J" "incseq ?F" "(\<Union>i. ?F i) = space ?J"
-      using F by auto
-    show "\<And>i. measure ?J (?F i) \<noteq> \<infinity>"
-      using F by (simp add: J.measure_times setprod_PInf)
-    have "measure_space (Pi\<^isub>M J M)" by default
-    then show "measure_space (sigma ?J)"
-      by (simp add: product_algebra_def sigma_def)
-    show "?R \<in> measure_preserving (Pi\<^isub>M K M) ?J"
-    proof (simp add: measure_preserving_def measurable_def product_algebra_generator_def del: vimage_Int,
-           safe intro!: restrict_extensional)
-      fix x k assume "k \<in> J" "x \<in> (\<Pi> i\<in>K. space (M i))"
-      then show "x k \<in> space (M k)" using `J \<subseteq> K` by auto
-    next
-      fix E assume "E \<in> (\<Pi> i\<in>J. sets (M i))"
-      then have E: "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)" by auto
-      then have *: "?R -` Pi\<^isub>E J E \<inter> (\<Pi>\<^isub>E i\<in>K. space (M i)) = (\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i))"
-        (is "?X = Pi\<^isub>E K ?M")
-        using `J \<subseteq> K` sets_into_space by (auto simp: Pi_iff split: split_if_asm) blast+
-      with E show "?X \<in> sets (Pi\<^isub>M K M)"
-        by (auto intro!: product_algebra_generatorI)
-      have "measure (Pi\<^isub>M J M) (Pi\<^isub>E J E) = (\<Prod>i\<in>J. measure (M i) (?M i))"
-        using E by (simp add: J.measure_times)
-      also have "\<dots> = measure (Pi\<^isub>M K M) ?X"
-        unfolding * using E `finite K` `J \<subseteq> K`
-        by (auto simp: K.measure_times M.measure_space_1
-                 cong del: setprod_cong
-                 intro!: setprod_mono_one_left)
-      finally show "measure (Pi\<^isub>M J M) (Pi\<^isub>E J E) = measure (Pi\<^isub>M K M) ?X" .
-    qed
-  qed
-  then show ?thesis
-    by (simp add: product_algebra_def sigma_def)
+  shows "(\<Pi>\<^isub>M i\<in>J. M i) = distr (\<Pi>\<^isub>M i\<in>K. M i) (\<Pi>\<^isub>M i\<in>J. M i) (\<lambda>f. restrict f J)" (is "?P = ?D")
+proof (rule measure_eqI_generator_eq)
+  have "finite J" using `J \<subseteq> K` `finite K` by (auto simp add: finite_subset)
+  interpret J: finite_product_prob_space M J proof qed fact
+  interpret K: finite_product_prob_space M K proof qed fact
+
+  let ?J = "{Pi\<^isub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
+  let ?F = "\<lambda>i. \<Pi>\<^isub>E k\<in>J. space (M k)"
+  let ?\<Omega> = "(\<Pi>\<^isub>E k\<in>J. space (M k))"
+  show "Int_stable ?J"
+    by (rule Int_stable_PiE)
+  show "range ?F \<subseteq> ?J" "incseq ?F" "(\<Union>i. ?F i) = ?\<Omega>"
+    using `finite J` by (auto intro!: prod_algebraI_finite)
+  { fix i show "emeasure ?P (?F i) \<noteq> \<infinity>" by simp }
+  show "?J \<subseteq> Pow ?\<Omega>" by (auto simp: Pi_iff dest: sets_into_space)
+  show "sets (\<Pi>\<^isub>M i\<in>J. M i) = sigma_sets ?\<Omega> ?J" "sets ?D = sigma_sets ?\<Omega> ?J"
+    using `finite J` by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff)
+  
+  fix X assume "X \<in> ?J"
+  then obtain E where [simp]: "X = Pi\<^isub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto
+  with `finite J` have X: "X \<in> sets (Pi\<^isub>M J M)" by auto
+
+  have "emeasure ?P X = (\<Prod> i\<in>J. emeasure (M i) (E i))"
+    using E by (simp add: J.measure_times)
+  also have "\<dots> = (\<Prod> i\<in>J. emeasure (M i) (if i \<in> J then E i else space (M i)))"
+    by simp
+  also have "\<dots> = (\<Prod> i\<in>K. emeasure (M i) (if i \<in> J then E i else space (M i)))"
+    using `finite K` `J \<subseteq> K`
+    by (intro setprod_mono_one_left) (auto simp: M.emeasure_space_1)
+  also have "\<dots> = emeasure (Pi\<^isub>M K M) (\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i))"
+    using E by (simp add: K.measure_times)
+  also have "(\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i)) = (\<lambda>f. restrict f J) -` Pi\<^isub>E J E \<inter> (\<Pi>\<^isub>E i\<in>K. space (M i))"
+    using `J \<subseteq> K` sets_into_space E by (force simp:  Pi_iff split: split_if_asm)
+  finally show "emeasure (Pi\<^isub>M J M) X = emeasure ?D X"
+    using X `J \<subseteq> K` apply (subst emeasure_distr)
+    by (auto intro!: measurable_restrict_subset simp: space_PiM)
 qed
 
-lemma (in product_prob_space) measurable_restrict:
-  assumes *: "J \<noteq> {}" "J \<subseteq> K" "finite K"
-  shows "(\<lambda>f. restrict f J) \<in> measurable (\<Pi>\<^isub>M i\<in>K. M i) (\<Pi>\<^isub>M i\<in>J. M i)"
-  using measure_preserving_restrict[OF *]
-  by (rule measure_preservingD2)
+abbreviation (in product_prob_space)
+  "emb L K X \<equiv> prod_emb L M K X"
+
+lemma (in product_prob_space) emeasure_prod_emb[simp]:
+  assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^isub>M J M)"
+  shows "emeasure (Pi\<^isub>M L M) (emb L J X) = emeasure (Pi\<^isub>M J M) X"
+  by (subst distr_restrict[OF L])
+     (simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X)
 
-definition (in product_prob_space)
-  "emb J K X = (\<lambda>x. restrict x K) -` X \<inter> space (Pi\<^isub>M J M)"
+lemma (in product_prob_space) prod_emb_injective:
+  assumes "J \<noteq> {}" "J \<subseteq> L" "finite J" and sets: "X \<in> sets (Pi\<^isub>M J M)" "Y \<in> sets (Pi\<^isub>M J M)"
+  assumes "prod_emb L M J X = prod_emb L M J Y"
+  shows "X = Y"
+proof (rule injective_vimage_restrict)
+  show "X \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))"
+    using sets[THEN sets_into_space] by (auto simp: space_PiM)
+  have "\<forall>i\<in>L. \<exists>x. x \<in> space (M i)"
+    using M.not_empty by auto
+  from bchoice[OF this]
+  show "(\<Pi>\<^isub>E i\<in>L. space (M i)) \<noteq> {}" by auto
+  show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i)) = (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i))"
+    using `prod_emb L M J X = prod_emb L M J Y` by (simp add: prod_emb_def)
+qed fact
 
-lemma (in product_prob_space) emb_trans[simp]:
-  "J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> emb L K (emb K J X) = emb L J X"
-  by (auto simp add: Int_absorb1 emb_def)
-
-lemma (in product_prob_space) emb_empty[simp]:
-  "emb K J {} = {}"
-  by (simp add: emb_def)
+definition (in product_prob_space) generator :: "('i \<Rightarrow> 'a) set set" where
+  "generator = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^isub>M J M))"
 
-lemma (in product_prob_space) emb_Pi:
-  assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
-  shows "emb K J (Pi\<^isub>E J X) = (\<Pi>\<^isub>E i\<in>K. if i \<in> J then X i else space (M i))"
-  using assms space_closed
-  by (auto simp: emb_def Pi_iff split: split_if_asm) blast+
+lemma (in product_prob_space) generatorI':
+  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> emb I J X \<in> generator"
+  unfolding generator_def by auto
 
-lemma (in product_prob_space) emb_injective:
-  assumes "J \<noteq> {}" "J \<subseteq> L" "finite J" and sets: "X \<in> sets (Pi\<^isub>M J M)" "Y \<in> sets (Pi\<^isub>M J M)"
-  assumes "emb L J X = emb L J Y"
-  shows "X = Y"
-proof -
-  interpret J: finite_product_sigma_finite M J by default fact
-  show "X = Y"
-  proof (rule injective_vimage_restrict)
-    show "X \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))"
-      using J.sets_into_space sets by auto
-    have "\<forall>i\<in>L. \<exists>x. x \<in> space (M i)"
-      using M.not_empty by auto
-    from bchoice[OF this]
-    show "(\<Pi>\<^isub>E i\<in>L. space (M i)) \<noteq> {}" by auto
-    show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i)) = (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i))"
-      using `emb L J X = emb L J Y` by (simp add: emb_def)
-  qed fact
+lemma (in product_prob_space) algebra_generator:
+  assumes "I \<noteq> {}" shows "algebra (\<Pi>\<^isub>E i\<in>I. space (M i)) generator" (is "algebra ?\<Omega> ?G")
+proof
+  let ?G = generator
+  show "?G \<subseteq> Pow ?\<Omega>"
+    by (auto simp: generator_def prod_emb_def)
+  from `I \<noteq> {}` obtain i where "i \<in> I" by auto
+  then show "{} \<in> ?G"
+    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\<lambda>i. {}"]
+             simp: sigma_sets.Empty generator_def prod_emb_def)
+  from `i \<in> I` show "?\<Omega> \<in> ?G"
+    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\<lambda>i. space (M i))"]
+             simp: generator_def prod_emb_def)
+  fix A assume "A \<in> ?G"
+  then obtain JA XA where XA: "JA \<noteq> {}" "finite JA" "JA \<subseteq> I" "XA \<in> sets (Pi\<^isub>M JA M)" and A: "A = emb I JA XA"
+    by (auto simp: generator_def)
+  fix B assume "B \<in> ?G"
+  then obtain JB XB where XB: "JB \<noteq> {}" "finite JB" "JB \<subseteq> I" "XB \<in> sets (Pi\<^isub>M JB M)" and B: "B = emb I JB XB"
+    by (auto simp: generator_def)
+  let ?RA = "emb (JA \<union> JB) JA XA"
+  let ?RB = "emb (JA \<union> JB) JB XB"
+  have *: "A - B = emb I (JA \<union> JB) (?RA - ?RB)" "A \<union> B = emb I (JA \<union> JB) (?RA \<union> ?RB)"
+    using XA A XB B by auto
+  show "A - B \<in> ?G" "A \<union> B \<in> ?G"
+    unfolding * using XA XB by (safe intro!: generatorI') auto
 qed
 
-lemma (in product_prob_space) emb_id:
-  "B \<subseteq> (\<Pi>\<^isub>E i\<in>L. space (M i)) \<Longrightarrow> emb L L B = B"
-  by (auto simp: emb_def Pi_iff subset_eq extensional_restrict)
-
-lemma (in product_prob_space) emb_simps:
-  shows "emb L K (A \<union> B) = emb L K A \<union> emb L K B"
-    and "emb L K (A \<inter> B) = emb L K A \<inter> emb L K B"
-    and "emb L K (A - B) = emb L K A - emb L K B"
-  by (auto simp: emb_def)
-
-lemma (in product_prob_space) measurable_emb[intro,simp]:
-  assumes *: "J \<noteq> {}" "J \<subseteq> L" "finite L" "X \<in> sets (Pi\<^isub>M J M)"
-  shows "emb L J X \<in> sets (Pi\<^isub>M L M)"
-  using measurable_restrict[THEN measurable_sets, OF *] by (simp add: emb_def)
-
-lemma (in product_prob_space) measure_emb[intro,simp]:
-  assumes *: "J \<noteq> {}" "J \<subseteq> L" "finite L" "X \<in> sets (Pi\<^isub>M J M)"
-  shows "measure (Pi\<^isub>M L M) (emb L J X) = measure (Pi\<^isub>M J M) X"
-  using measure_preserving_restrict[THEN measure_preservingD, OF *]
-  by (simp add: emb_def)
-
-definition (in product_prob_space) generator :: "('i \<Rightarrow> 'a) measure_space" where
-  "generator = \<lparr>
-    space = (\<Pi>\<^isub>E i\<in>I. space (M i)),
-    sets = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^isub>M J M)),
-    measure = undefined
-  \<rparr>"
+lemma (in product_prob_space) sets_PiM_generator:
+  assumes "I \<noteq> {}" shows "sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
+proof
+  show "sets (Pi\<^isub>M I M) \<subseteq> sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
+    unfolding sets_PiM
+  proof (safe intro!: sigma_sets_subseteq)
+    fix A assume "A \<in> prod_algebra I M" with `I \<noteq> {}` show "A \<in> generator"
+      by (auto intro!: generatorI' elim!: prod_algebraE)
+  qed
+qed (auto simp: generator_def space_PiM[symmetric] intro!: sigma_sets_subset)
 
 lemma (in product_prob_space) generatorI:
-  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> sets generator"
-  unfolding generator_def by auto
-
-lemma (in product_prob_space) generatorI':
-  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> emb I J X \<in> sets generator"
+  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> generator"
   unfolding generator_def by auto
 
-lemma (in product_sigma_finite)
-  assumes "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
-  shows measure_fold_integral:
-    "measure (Pi\<^isub>M (I \<union> J) M) A = (\<integral>\<^isup>+x. measure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M)) \<partial>Pi\<^isub>M I M)" (is ?I)
-    and measure_fold_measurable:
-    "(\<lambda>x. measure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M))) \<in> borel_measurable (Pi\<^isub>M I M)" (is ?B)
-proof -
-  interpret I: finite_product_sigma_finite M I by default fact
-  interpret J: finite_product_sigma_finite M J by default fact
-  interpret IJ: pair_sigma_finite I.P J.P ..
-  show ?I
-    unfolding measure_fold[OF assms]
-    apply (subst IJ.pair_measure_alt)
-    apply (intro measurable_sets[OF _ A] measurable_merge assms)
-    apply (auto simp: vimage_compose[symmetric] comp_def space_pair_measure
-      intro!: I.positive_integral_cong)
-    done
-
-  have "(\<lambda>(x, y). merge I x J y) -` A \<inter> space (I.P \<Otimes>\<^isub>M J.P) \<in> sets (I.P \<Otimes>\<^isub>M J.P)"
-    by (intro measurable_sets[OF _ A] measurable_merge assms)
-  from IJ.measure_cut_measurable_fst[OF this]
-  show ?B
-    apply (auto simp: vimage_compose[symmetric] comp_def space_pair_measure)
-    apply (subst (asm) measurable_cong)
-    apply auto
-    done
-qed
-
 definition (in product_prob_space)
   "\<mu>G A =
-    (THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^isub>M J M). A = emb I J X \<longrightarrow> x = measure (Pi\<^isub>M J M) X))"
+    (THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^isub>M J M). A = emb I J X \<longrightarrow> x = emeasure (Pi\<^isub>M J M) X))"
 
 lemma (in product_prob_space) \<mu>G_spec:
   assumes J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
-  shows "\<mu>G A = measure (Pi\<^isub>M J M) X"
+  shows "\<mu>G A = emeasure (Pi\<^isub>M J M) X"
   unfolding \<mu>G_def
 proof (intro the_equality allI impI ballI)
   fix K Y assume K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "A = emb I K Y" "Y \<in> sets (Pi\<^isub>M K M)"
-  have "measure (Pi\<^isub>M K M) Y = measure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) K Y)"
+  have "emeasure (Pi\<^isub>M K M) Y = emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) K Y)"
     using K J by simp
   also have "emb (K \<union> J) K Y = emb (K \<union> J) J X"
-    using K J by (simp add: emb_injective[of "K \<union> J" I])
-  also have "measure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) J X) = measure (Pi\<^isub>M J M) X"
+    using K J by (simp add: prod_emb_injective[of "K \<union> J" I])
+  also have "emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) J X) = emeasure (Pi\<^isub>M J M) X"
     using K J by simp
-  finally show "measure (Pi\<^isub>M J M) X = measure (Pi\<^isub>M K M) Y" ..
+  finally show "emeasure (Pi\<^isub>M J M) X = emeasure (Pi\<^isub>M K M) Y" ..
 qed (insert J, force)
 
 lemma (in product_prob_space) \<mu>G_eq:
-  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = measure (Pi\<^isub>M J M) X"
+  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = emeasure (Pi\<^isub>M J M) X"
   by (intro \<mu>G_spec) auto
 
 lemma (in product_prob_space) generator_Ex:
-  assumes *: "A \<in> sets generator"
-  shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A = emb I J X \<and> \<mu>G A = measure (Pi\<^isub>M J M) X"
+  assumes *: "A \<in> generator"
+  shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A = emb I J X \<and> \<mu>G A = emeasure (Pi\<^isub>M J M) X"
 proof -
   from * obtain J X where J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
     unfolding generator_def by auto
@@ -231,11 +260,11 @@
 qed
 
 lemma (in product_prob_space) generatorE:
-  assumes A: "A \<in> sets generator"
-  obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A" "\<mu>G A = measure (Pi\<^isub>M J M) X"
+  assumes A: "A \<in> generator"
+  obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A" "\<mu>G A = emeasure (Pi\<^isub>M J M) X"
 proof -
   from generator_Ex[OF A] obtain X J where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A"
-    "\<mu>G A = measure (Pi\<^isub>M J M) X" by auto
+    "\<mu>G A = emeasure (Pi\<^isub>M J M) X" by auto
   then show thesis by (intro that) auto
 qed
 
@@ -243,11 +272,7 @@
   assumes "finite J" "finite K" "J \<inter> K = {}" and A: "A \<in> sets (Pi\<^isub>M (J \<union> K) M)" and x: "x \<in> space (Pi\<^isub>M J M)"
   shows "merge J x K -` A \<inter> space (Pi\<^isub>M K M) \<in> sets (Pi\<^isub>M K M)"
 proof -
-  interpret J: finite_product_sigma_algebra M J by default fact
-  interpret K: finite_product_sigma_algebra M K by default fact
-  interpret JK: pair_sigma_algebra J.P K.P ..
-
-  from JK.measurable_cut_fst[OF
+  from sets_Pair1[OF
     measurable_merge[THEN measurable_sets, OF `J \<inter> K = {}`], OF A, of x] x
   show ?thesis
       by (simp add: space_pair_measure comp_def vimage_compose[symmetric])
@@ -266,75 +291,27 @@
   have [simp]: "(K - J) \<inter> (K \<union> J) = K - J" by auto
   have [simp]: "(K - J) \<inter> K = K - J" by auto
   from y `K \<subseteq> I` `J \<subseteq> I` show ?thesis
-    by (simp split: split_merge add: emb_def Pi_iff extensional_merge_sub set_eq_iff) auto
-qed
-
-definition (in product_prob_space) infprod_algebra :: "('i \<Rightarrow> 'a) measure_space" where
-  "infprod_algebra = sigma generator \<lparr> measure :=
-    (SOME \<mu>. (\<forall>s\<in>sets generator. \<mu> s = \<mu>G s) \<and>
-       prob_space \<lparr>space = space generator, sets = sets (sigma generator), measure = \<mu>\<rparr>)\<rparr>"
-
-syntax
-  "_PiP"  :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme"  ("(3PIP _:_./ _)" 10)
-
-syntax (xsymbols)
-  "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme"  ("(3\<Pi>\<^isub>P _\<in>_./ _)"   10)
-
-syntax (HTML output)
-  "_PiP" :: "[pttrn, 'i set, ('b, 'd) measure_space_scheme] => ('i => 'b, 'd) measure_space_scheme"  ("(3\<Pi>\<^isub>P _\<in>_./ _)"   10)
-
-abbreviation
-  "Pi\<^isub>P I M \<equiv> product_prob_space.infprod_algebra M I"
-
-translations
-  "PIP x:I. M" == "CONST Pi\<^isub>P I (%x. M)"
-
-lemma (in product_prob_space) algebra_generator:
-  assumes "I \<noteq> {}" shows "algebra generator"
-proof
-  let ?G = generator
-  show "sets ?G \<subseteq> Pow (space ?G)"
-    by (auto simp: generator_def emb_def)
-  from `I \<noteq> {}` obtain i where "i \<in> I" by auto
-  then show "{} \<in> sets ?G"
-    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\<lambda>i. {}"]
-      simp: product_algebra_def sigma_def sigma_sets.Empty generator_def emb_def)
-  from `i \<in> I` show "space ?G \<in> sets ?G"
-    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\<lambda>i. space (M i))"]
-      simp: generator_def emb_def)
-  fix A assume "A \<in> sets ?G"
-  then obtain JA XA where XA: "JA \<noteq> {}" "finite JA" "JA \<subseteq> I" "XA \<in> sets (Pi\<^isub>M JA M)" and A: "A = emb I JA XA"
-    by (auto simp: generator_def)
-  fix B assume "B \<in> sets ?G"
-  then obtain JB XB where XB: "JB \<noteq> {}" "finite JB" "JB \<subseteq> I" "XB \<in> sets (Pi\<^isub>M JB M)" and B: "B = emb I JB XB"
-    by (auto simp: generator_def)
-  let ?RA = "emb (JA \<union> JB) JA XA"
-  let ?RB = "emb (JA \<union> JB) JB XB"
-  interpret JAB: finite_product_sigma_algebra M "JA \<union> JB"
-    by default (insert XA XB, auto)
-  have *: "A - B = emb I (JA \<union> JB) (?RA - ?RB)" "A \<union> B = emb I (JA \<union> JB) (?RA \<union> ?RB)"
-    using XA A XB B by (auto simp: emb_simps)
-  then show "A - B \<in> sets ?G" "A \<union> B \<in> sets ?G"
-    using XA XB by (auto intro!: generatorI')
+    by (simp split: split_merge add: prod_emb_def Pi_iff extensional_merge_sub set_eq_iff space_PiM)
+       auto
 qed
 
 lemma (in product_prob_space) positive_\<mu>G: 
   assumes "I \<noteq> {}"
   shows "positive generator \<mu>G"
 proof -
-  interpret G!: algebra generator by (rule algebra_generator) fact
+  interpret G!: algebra "\<Pi>\<^isub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
   show ?thesis
   proof (intro positive_def[THEN iffD2] conjI ballI)
     from generatorE[OF G.empty_sets] guess J X . note this[simp]
     interpret J: finite_product_sigma_finite M J by default fact
     have "X = {}"
-      by (rule emb_injective[of J I]) simp_all
+      by (rule prod_emb_injective[of J I]) simp_all
     then show "\<mu>G {} = 0" by simp
   next
-    fix A assume "A \<in> sets generator"
+    fix A assume "A \<in> generator"
     from generatorE[OF this] guess J X . note this[simp]
     interpret J: finite_product_sigma_finite M J by default fact
-    show "0 \<le> \<mu>G A" by simp
+    show "0 \<le> \<mu>G A" by (simp add: emeasure_nonneg)
   qed
 qed
 
@@ -342,102 +319,47 @@
   assumes "I \<noteq> {}"
   shows "additive generator \<mu>G"
 proof -
-  interpret G!: algebra generator by (rule algebra_generator) fact
+  interpret G!: algebra "\<Pi>\<^isub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
   show ?thesis
   proof (intro additive_def[THEN iffD2] ballI impI)
-    fix A assume "A \<in> sets generator" with generatorE guess J X . note J = this
-    fix B assume "B \<in> sets generator" with generatorE guess K Y . note K = this
+    fix A assume "A \<in> generator" with generatorE guess J X . note J = this
+    fix B assume "B \<in> generator" with generatorE guess K Y . note K = this
     assume "A \<inter> B = {}"
     have JK: "J \<union> K \<noteq> {}" "J \<union> K \<subseteq> I" "finite (J \<union> K)"
       using J K by auto
     interpret JK: finite_product_sigma_finite M "J \<union> K" by default fact
     have JK_disj: "emb (J \<union> K) J X \<inter> emb (J \<union> K) K Y = {}"
-      apply (rule emb_injective[of "J \<union> K" I])
+      apply (rule prod_emb_injective[of "J \<union> K" I])
       apply (insert `A \<inter> B = {}` JK J K)
-      apply (simp_all add: JK.Int emb_simps)
+      apply (simp_all add: Int prod_emb_Int)
       done
     have AB: "A = emb I (J \<union> K) (emb (J \<union> K) J X)" "B = emb I (J \<union> K) (emb (J \<union> K) K Y)"
       using J K by simp_all
     then have "\<mu>G (A \<union> B) = \<mu>G (emb I (J \<union> K) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y))"
-      by (simp add: emb_simps)
-    also have "\<dots> = measure (Pi\<^isub>M (J \<union> K) M) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)"
-      using JK J(1, 4) K(1, 4) by (simp add: \<mu>G_eq JK.Un)
+      by simp
+    also have "\<dots> = emeasure (Pi\<^isub>M (J \<union> K) M) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)"
+      using JK J(1, 4) K(1, 4) by (simp add: \<mu>G_eq Un del: prod_emb_Un)
     also have "\<dots> = \<mu>G A + \<mu>G B"
-      using J K JK_disj by (simp add: JK.measure_additive[symmetric])
+      using J K JK_disj by (simp add: plus_emeasure[symmetric])
     finally show "\<mu>G (A \<union> B) = \<mu>G A + \<mu>G B" .
   qed
 qed
 
-lemma (in product_prob_space) finite_index_eq_finite_product:
-  assumes "finite I"
-  shows "sets (sigma generator) = sets (Pi\<^isub>M I M)"
-proof safe
-  interpret I: finite_product_sigma_algebra M I by default fact
-  have space_generator[simp]: "space generator = space (Pi\<^isub>M I M)"
-    by (simp add: generator_def product_algebra_def)
-  { fix A assume "A \<in> sets (sigma generator)"
-    then show "A \<in> sets I.P" unfolding sets_sigma
-    proof induct
-      case (Basic A)
-      from generatorE[OF this] guess J X . note J = this
-      with `finite I` have "emb I J X \<in> sets I.P" by auto
-      with `emb I J X = A` show "A \<in> sets I.P" by simp
-    qed auto }
-  { fix A assume A: "A \<in> sets I.P"
-    show "A \<in> sets (sigma generator)"
-    proof cases
-      assume "I = {}"
-      with I.P_empty[OF this] A
-      have "A = space generator \<or> A = {}" 
-        unfolding space_generator by auto
-      then show ?thesis
-        by (auto simp: sets_sigma simp del: space_generator
-                 intro: sigma_sets.Empty sigma_sets_top)
-    next
-      assume "I \<noteq> {}"
-      note A this
-      moreover with I.sets_into_space have "emb I I A = A" by (intro emb_id) auto
-      ultimately show "A \<in> sets (sigma generator)"
-        using `finite I` unfolding sets_sigma
-        by (intro sigma_sets.Basic generatorI[of I A]) auto
-  qed }
-qed
-
-lemma (in product_prob_space) extend_\<mu>G:
-  "\<exists>\<mu>. (\<forall>s\<in>sets generator. \<mu> s = \<mu>G s) \<and>
-       prob_space \<lparr>space = space generator, sets = sets (sigma generator), measure = \<mu>\<rparr>"
+lemma (in product_prob_space) emeasure_PiM_emb_not_empty:
+  assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. X i \<in> sets (M i)"
+  shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)"
 proof cases
-  assume "finite I"
-  interpret I: finite_product_prob_space M I by default fact
-  show ?thesis
-  proof (intro exI[of _ "measure (Pi\<^isub>M I M)"] ballI conjI)
-    fix A assume "A \<in> sets generator"
-    from generatorE[OF this] guess J X . note J = this
-    from J(1-4) `finite I` show "measure I.P A = \<mu>G A"
-      unfolding J(6)
-      by (subst J(5)[symmetric]) (simp add: measure_emb)
-  next
-    have [simp]: "space generator = space (Pi\<^isub>M I M)"
-      by (simp add: generator_def product_algebra_def)
-    have "\<lparr>space = space generator, sets = sets (sigma generator), measure = measure I.P\<rparr>
-      = I.P" (is "?P = _")
-      by (auto intro!: measure_space.equality simp: finite_index_eq_finite_product[OF `finite I`])
-    show "prob_space ?P"
-    proof
-      show "measure_space ?P" using `?P = I.P` by simp default
-      show "measure ?P (space ?P) = 1"
-        using I.measure_space_1 by simp
-    qed
-  qed
+  assume "finite I" with X show ?thesis by simp
 next
+  let ?\<Omega> = "\<Pi>\<^isub>E i\<in>I. space (M i)"
   let ?G = generator
   assume "\<not> finite I"
   then have I_not_empty: "I \<noteq> {}" by auto
-  interpret G!: algebra generator by (rule algebra_generator) fact
+  interpret G!: algebra ?\<Omega> generator by (rule algebra_generator) fact
   note \<mu>G_mono =
     G.additive_increasing[OF positive_\<mu>G[OF I_not_empty] additive_\<mu>G[OF I_not_empty], THEN increasingD]
 
-  { fix Z J assume J: "J \<noteq> {}" "finite J" "J \<subseteq> I" and Z: "Z \<in> sets ?G"
+  { fix Z J assume J: "J \<noteq> {}" "finite J" "J \<subseteq> I" and Z: "Z \<in> ?G"
 
     from `infinite I` `finite J` obtain k where k: "k \<in> I" "k \<notin> J"
       by (metis rev_finite_subset subsetI)
@@ -445,7 +367,7 @@
     moreover def K \<equiv> "insert k K'"
     moreover def X \<equiv> "emb K K' X'"
     ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^isub>M K M)" "Z = emb I K X"
-      "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = measure (Pi\<^isub>M K M) X"
+      "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^isub>M K M) X"
       by (auto simp: subset_insertI)
 
     let ?M = "\<lambda>y. merge J y (K - J) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M)"
@@ -455,9 +377,9 @@
       have **: "?M y \<in> sets (Pi\<^isub>M (K - J) M)"
         using J K y by (intro merge_sets) auto
       ultimately
-      have ***: "(merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<in> sets ?G"
+      have ***: "(merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<in> ?G"
         using J K by (intro generatorI) auto
-      have "\<mu>G (merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = measure (Pi\<^isub>M (K - J) M) (?M y)"
+      have "\<mu>G (merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = emeasure (Pi\<^isub>M (K - J) M) (?M y)"
         unfolding * using K J by (subst \<mu>G_eq[OF _ _ _ **]) auto
       note * ** *** this }
     note merge_in_G = this
@@ -467,16 +389,16 @@
     interpret J: finite_product_prob_space M J by default fact+
     interpret KmJ: finite_product_prob_space M "K - J" by default fact+
 
-    have "\<mu>G Z = measure (Pi\<^isub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)"
+    have "\<mu>G Z = emeasure (Pi\<^isub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)"
       using K J by simp
-    also have "\<dots> = (\<integral>\<^isup>+ x. measure (Pi\<^isub>M (K - J) M) (?M x) \<partial>Pi\<^isub>M J M)"
-      using K J by (subst measure_fold_integral) auto
+    also have "\<dots> = (\<integral>\<^isup>+ x. emeasure (Pi\<^isub>M (K - J) M) (?M x) \<partial>Pi\<^isub>M J M)"
+      using K J by (subst emeasure_fold_integral) auto
     also have "\<dots> = (\<integral>\<^isup>+ y. \<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<partial>Pi\<^isub>M J M)"
       (is "_ = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)")
-    proof (intro J.positive_integral_cong)
+    proof (intro positive_integral_cong)
       fix x assume x: "x \<in> space (Pi\<^isub>M J M)"
       with K merge_in_G(2)[OF this]
-      show "measure (Pi\<^isub>M (K - J) M) (?M x) = \<mu>G (?MZ x)"
+      show "emeasure (Pi\<^isub>M (K - J) M) (?M x) = \<mu>G (?MZ x)"
         unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst \<mu>G_eq) auto
     qed
     finally have fold: "\<mu>G Z = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)" .
@@ -490,21 +412,18 @@
     let ?q = "\<lambda>y. \<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M))"
     have "?q \<in> borel_measurable (Pi\<^isub>M J M)"
       unfolding `Z = emb I K X` using J K merge_in_G(3)
-      by (simp add: merge_in_G  \<mu>G_eq measure_fold_measurable
-               del: space_product_algebra cong: measurable_cong)
+      by (simp add: merge_in_G  \<mu>G_eq emeasure_fold_measurable cong: measurable_cong)
     note this fold le_1 merge_in_G(3) }
   note fold = this
 
-  have "\<exists>\<mu>. (\<forall>s\<in>sets ?G. \<mu> s = \<mu>G s) \<and>
-    measure_space \<lparr>space = space ?G, sets = sets (sigma ?G), measure = \<mu>\<rparr>"
-    (is "\<exists>\<mu>. _ \<and> measure_space (?ms \<mu>)")
+  have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>"
   proof (rule G.caratheodory_empty_continuous[OF positive_\<mu>G additive_\<mu>G])
-    fix A assume "A \<in> sets ?G"
+    fix A assume "A \<in> ?G"
     with generatorE guess J X . note JX = this
     interpret JK: finite_product_prob_space M J by default fact+
     from JX show "\<mu>G A \<noteq> \<infinity>" by simp
   next
-    fix A assume A: "range A \<subseteq> sets ?G" "decseq A" "(\<Inter>i. A i) = {}"
+    fix A assume A: "range A \<subseteq> ?G" "decseq A" "(\<Inter>i. A i) = {}"
     then have "decseq (\<lambda>i. \<mu>G (A i))"
       by (auto intro!: \<mu>G_mono simp: decseq_def)
     moreover
@@ -515,7 +434,7 @@
         using A positive_\<mu>G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def)
       ultimately have "0 < ?a" by auto
 
-      have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = measure (Pi\<^isub>M J M) X"
+      have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (Pi\<^isub>M J M) X"
         using A by (intro allI generator_Ex) auto
       then obtain J' X' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I" "\<And>n. X' n \<in> sets (Pi\<^isub>M (J' n) M)"
         and A': "\<And>n. A n = emb I (J' n) (X' n)"
@@ -524,8 +443,8 @@
       moreover def X \<equiv> "\<lambda>n. emb (J n) (J' n) (X' n)"
       ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. X n \<in> sets (Pi\<^isub>M (J n) M)"
         by auto
-      with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> sets ?G"
-        unfolding J_def X_def by (subst emb_trans) (insert A, auto)
+      with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> ?G"
+        unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto)
 
       have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
         unfolding J_def by force
@@ -538,8 +457,8 @@
 
       let ?M = "\<lambda>K Z y. merge K y (I - K) -` Z \<inter> space (Pi\<^isub>M I M)"
 
-      { fix Z k assume Z: "range Z \<subseteq> sets ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)"
-        then have Z_sets: "\<And>n. Z n \<in> sets ?G" by auto
+      { fix Z k assume Z: "range Z \<subseteq> ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)"
+        then have Z_sets: "\<And>n. Z n \<in> ?G" by auto
         fix J' assume J': "J' \<noteq> {}" "finite J'" "J' \<subseteq> I"
         interpret J': finite_product_prob_space M J' by default fact+
 
@@ -552,13 +471,13 @@
             by (rule measurable_sets) auto }
         note Q_sets = this
 
-        have "?a / 2^(k+1) \<le> (INF n. measure (Pi\<^isub>M J' M) (?Q n))"
+        have "?a / 2^(k+1) \<le> (INF n. emeasure (Pi\<^isub>M J' M) (?Q n))"
         proof (intro INF_greatest)
           fix n
           have "?a / 2^k \<le> \<mu>G (Z n)" using Z by auto
           also have "\<dots> \<le> (\<integral>\<^isup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^isub>M J' M)"
-            unfolding fold(2)[OF J' `Z n \<in> sets ?G`]
-          proof (intro J'.positive_integral_mono)
+            unfolding fold(2)[OF J' `Z n \<in> ?G`]
+          proof (intro positive_integral_mono)
             fix x assume x: "x \<in> space (Pi\<^isub>M J' M)"
             then have "?q n x \<le> 1 + 0"
               using J' Z fold(3) Z_sets by auto
@@ -568,15 +487,15 @@
             with x show "?q n x \<le> indicator (?Q n) x + ?a / 2^(k+1)"
               by (auto split: split_indicator simp del: power_Suc)
           qed
-          also have "\<dots> = measure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)"
-            using `0 \<le> ?a` Q_sets J'.measure_space_1
-            by (subst J'.positive_integral_add) auto
-          finally show "?a / 2^(k+1) \<le> measure (Pi\<^isub>M J' M) (?Q n)" using `?a \<le> 1`
-            by (cases rule: ereal2_cases[of ?a "measure (Pi\<^isub>M J' M) (?Q n)"])
+          also have "\<dots> = emeasure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)"
+            using `0 \<le> ?a` Q_sets J'.emeasure_space_1
+            by (subst positive_integral_add) auto
+          finally show "?a / 2^(k+1) \<le> emeasure (Pi\<^isub>M J' M) (?Q n)" using `?a \<le> 1`
+            by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^isub>M J' M) (?Q n)"])
                (auto simp: field_simps)
         qed
-        also have "\<dots> = measure (Pi\<^isub>M J' M) (\<Inter>n. ?Q n)"
-        proof (intro J'.continuity_from_above)
+        also have "\<dots> = emeasure (Pi\<^isub>M J' M) (\<Inter>n. ?Q n)"
+        proof (intro INF_emeasure_decseq)
           show "range ?Q \<subseteq> sets (Pi\<^isub>M J' M)" using Q_sets by auto
           show "decseq ?Q"
             unfolding decseq_def
@@ -587,13 +506,13 @@
             also have "?q n x \<le> ?q m x"
             proof (rule \<mu>G_mono)
               from fold(4)[OF J', OF Z_sets x]
-              show "?M J' (Z n) x \<in> sets ?G" "?M J' (Z m) x \<in> sets ?G" by auto
+              show "?M J' (Z n) x \<in> ?G" "?M J' (Z m) x \<in> ?G" by auto
               show "?M J' (Z n) x \<subseteq> ?M J' (Z m) x"
                 using `decseq Z`[THEN decseqD, OF `m \<le> n`] by auto
             qed
             finally show "?a / 2^(k+1) \<le> ?q m x" .
           qed
-        qed (intro J'.finite_measure Q_sets)
+        qed simp
         finally have "(\<Inter>n. ?Q n) \<noteq> {}"
           using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
         then have "\<exists>w\<in>space (Pi\<^isub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto }
@@ -631,12 +550,12 @@
               show "w k \<in> space (Pi\<^isub>M (J (Suc k)) M)"
                 using Suc by simp
               then show "restrict (w k) (J k) = w k"
-                by (simp add: extensional_restrict)
+                by (simp add: extensional_restrict space_PiM)
             qed
           next
             assume "J k \<noteq> J (Suc k)"
             with J_mono[of k "Suc k"] have "J (Suc k) - J k \<noteq> {}" (is "?D \<noteq> {}") by auto
-            have "range (\<lambda>n. ?M (J k) (A n) (w k)) \<subseteq> sets ?G"
+            have "range (\<lambda>n. ?M (J k) (A n) (w k)) \<subseteq> ?G"
               "decseq (\<lambda>n. ?M (J k) (A n) (w k))"
               "\<forall>n. ?a / 2 ^ (k + 1) \<le> \<mu>G (?M (J k) (A n) (w k))"
               using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc
@@ -651,11 +570,11 @@
               by (auto intro!: ext split: split_merge)
             have *: "\<And>n. ?M ?D (?M (J k) (A n) (w k)) w' = ?M (J (Suc k)) (A n) ?w"
               using w'(1) J(3)[of "Suc k"]
-              by (auto split: split_merge intro!: extensional_merge_sub) force+
+              by (auto simp: space_PiM split: split_merge intro!: extensional_merge_sub) force+
             show ?thesis
               apply (rule exI[of _ ?w])
               using w' J_mono[of k "Suc k"] wk unfolding *
-              apply (auto split: split_merge intro!: extensional_merge_sub ext)
+              apply (auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM)
               apply (force simp: extensional_def)
               done
           qed
@@ -675,7 +594,7 @@
         then have "merge (J k) (w k) (I - J k) x \<in> A k" by auto
         then have "\<exists>x\<in>A k. restrict x (J k) = w k"
           using `w k \<in> space (Pi\<^isub>M (J k) M)`
-          by (intro rev_bexI) (auto intro!: ext simp: extensional_def)
+          by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM)
         ultimately have "w k \<in> space (Pi\<^isub>M (J k) M)"
           "\<exists>x\<in>A k. restrict x (J k) = w k"
           "k \<noteq> 0 \<Longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1)"
@@ -707,17 +626,17 @@
       have w'_simps2: "\<And>i. i \<notin> (\<Union>k. J k) \<Longrightarrow> i \<in> I \<Longrightarrow> w' i \<in> space (M i)"
         using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]])
       { fix i assume "i \<in> I" then have "w' i \<in> space (M i)"
-          using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq)+ }
+          using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ }
       note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this
 
       have w': "w' \<in> space (Pi\<^isub>M I M)"
-        using w(1) by (auto simp add: Pi_iff extensional_def)
+        using w(1) by (auto simp add: Pi_iff extensional_def space_PiM)
 
       { fix n
         have "restrict w' (J n) = w n" using w(1)
-          by (auto simp add: fun_eq_iff restrict_def Pi_iff extensional_def)
+          by (auto simp add: fun_eq_iff restrict_def Pi_iff extensional_def space_PiM)
         with w[of n] obtain x where "x \<in> A n" "restrict x (J n) = restrict w' (J n)" by auto
-        then have "w' \<in> A n" unfolding A_eq using w' by (auto simp: emb_def) }
+        then have "w' \<in> A n" unfolding A_eq using w' by (auto simp: prod_emb_def space_PiM) }
       then have "w' \<in> (\<Inter>i. A i)" by auto
       with `(\<Inter>i. A i) = {}` show False by auto
     qed
@@ -726,276 +645,76 @@
   qed fact+
   then guess \<mu> .. note \<mu> = this
   show ?thesis
-  proof (intro exI[of _ \<mu>] conjI)
-    show "\<forall>S\<in>sets ?G. \<mu> S = \<mu>G S" using \<mu> by simp
-    show "prob_space (?ms \<mu>)"
-    proof
-      show "measure_space (?ms \<mu>)" using \<mu> by simp
-      obtain i where "i \<in> I" using I_not_empty by auto
-      interpret i: finite_product_sigma_finite M "{i}" by default auto
-      let ?X = "\<Pi>\<^isub>E i\<in>{i}. space (M i)"
-      have X: "?X \<in> sets (Pi\<^isub>M {i} M)"
-        by auto
-      with `i \<in> I` have "emb I {i} ?X \<in> sets generator"
-        by (intro generatorI') auto
-      with \<mu> have "\<mu> (emb I {i} ?X) = \<mu>G (emb I {i} ?X)" by auto
-      with \<mu>G_eq[OF _ _ _ X] `i \<in> I` 
-      have "\<mu> (emb I {i} ?X) = measure (M i) (space (M i))"
-        by (simp add: i.measure_times)
-      also have "emb I {i} ?X = space (Pi\<^isub>P I M)"
-        using `i \<in> I` by (auto simp: emb_def infprod_algebra_def generator_def)
-      finally show "measure (?ms \<mu>) (space (?ms \<mu>)) = 1"
-        using M.measure_space_1 by (simp add: infprod_algebra_def)
-    qed
+  proof (subst emeasure_extend_measure_Pair[OF PiM_def, of I M \<mu> J X])
+    from assms show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))"
+      by (simp add: Pi_iff)
+  next
+    fix J X assume J: "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))"
+    then show "emb I J (Pi\<^isub>E J X) \<in> Pow (\<Pi>\<^isub>E i\<in>I. space (M i))"
+      by (auto simp: Pi_iff prod_emb_def dest: sets_into_space)
+    have "emb I J (Pi\<^isub>E J X) \<in> generator"
+      using J `I \<noteq> {}` by (intro generatorI') auto
+    then have "\<mu> (emb I J (Pi\<^isub>E J X)) = \<mu>G (emb I J (Pi\<^isub>E J X))"
+      using \<mu> by simp
+    also have "\<dots> = (\<Prod> j\<in>J. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
+      using J  `I \<noteq> {}` by (subst \<mu>G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff)
+    also have "\<dots> = (\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}.
+      if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
+      using J `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
+    finally show "\<mu> (emb I J (Pi\<^isub>E J X)) = \<dots>" .
+  next
+    let ?F = "\<lambda>j. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j))"
+    have "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) = (\<Prod>j\<in>J. ?F j)"
+      using X `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
+    then show "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) =
+      emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)"
+      using X by (auto simp add: emeasure_PiM) 
+  next
+    show "positive (sets (Pi\<^isub>M I M)) \<mu>" "countably_additive (sets (Pi\<^isub>M I M)) \<mu>"
+      using \<mu> unfolding sets_PiM_generator[OF `I \<noteq> {}`] by (auto simp: measure_space_def)
   qed
 qed
 
-lemma (in product_prob_space) infprod_spec:
-  "(\<forall>s\<in>sets generator. measure (Pi\<^isub>P I M) s = \<mu>G s) \<and> prob_space (Pi\<^isub>P I M)"
-  (is "?Q infprod_algebra")
-  unfolding infprod_algebra_def
-  by (rule someI2_ex[OF extend_\<mu>G])
-     (auto simp: sigma_def generator_def)
-
-sublocale product_prob_space \<subseteq> P: prob_space "Pi\<^isub>P I M"
-  using infprod_spec by simp
-
-lemma (in product_prob_space) measure_infprod_emb:
-  assumes "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)"
-  shows "\<mu> (emb I J X) = measure (Pi\<^isub>M J M) X"
-proof -
-  have "emb I J X \<in> sets generator"
-    using assms by (rule generatorI')
-  with \<mu>G_eq[OF assms] infprod_spec show ?thesis by auto
-qed
-
-lemma (in product_prob_space) measurable_component:
-  assumes "i \<in> I"
-  shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>P I M) (M i)"
-proof (unfold measurable_def, safe)
-  fix x assume "x \<in> space (Pi\<^isub>P I M)"
-  then show "x i \<in> space (M i)"
-    using `i \<in> I` by (auto simp: infprod_algebra_def generator_def)
-next
-  fix A assume "A \<in> sets (M i)"
-  with `i \<in> I` have
-    "(\<Pi>\<^isub>E x \<in> {i}. A) \<in> sets (Pi\<^isub>M {i} M)"
-    "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M) = emb I {i} (\<Pi>\<^isub>E x \<in> {i}. A)"
-    by (auto simp: infprod_algebra_def generator_def emb_def)
-  from generatorI[OF _ _ _ this] `i \<in> I`
-  show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M) \<in> sets (Pi\<^isub>P I M)"
-    unfolding infprod_algebra_def by auto
-qed
-
-lemma (in product_prob_space) emb_in_infprod_algebra[intro]:
-  fixes J assumes J: "finite J" "J \<subseteq> I" and X: "X \<in> sets (Pi\<^isub>M J M)"
-  shows "emb I J X \<in> sets (\<Pi>\<^isub>P i\<in>I. M i)"
-proof cases
-  assume "J = {}"
-  with X have "emb I J X = space (\<Pi>\<^isub>P i\<in>I. M i) \<or> emb I J X = {}"
-    by (auto simp: emb_def infprod_algebra_def generator_def
-                   product_algebra_def product_algebra_generator_def image_constant sigma_def)
-  then show ?thesis by auto
-next
-  assume "J \<noteq> {}"
-  show ?thesis unfolding infprod_algebra_def
-    by simp (intro in_sigma generatorI'  `J \<noteq> {}` J X)
-qed
-
-lemma (in product_prob_space) finite_measure_infprod_emb:
-  assumes "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)"
-  shows "\<mu>' (emb I J X) = finite_measure.\<mu>' (Pi\<^isub>M J M) X"
-proof -
-  interpret J: finite_product_prob_space M J by default fact+
-  from assms have "emb I J X \<in> sets (Pi\<^isub>P I M)" by auto
-  with assms show "\<mu>' (emb I J X) = J.\<mu>' X"
-    unfolding \<mu>'_def J.\<mu>'_def
-    unfolding measure_infprod_emb[OF assms]
-    by auto
-qed
-
-lemma (in finite_product_prob_space) finite_measure_times:
-  assumes "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)"
-  shows "\<mu>' (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu>' i (A i))"
-  using assms
-  unfolding \<mu>'_def M.\<mu>'_def
-  by (subst measure_times[OF assms])
-     (auto simp: finite_measure_eq M.finite_measure_eq setprod_ereal)
-
-lemma (in product_prob_space) finite_measure_infprod_emb_Pi:
-  assumes J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> X j \<in> sets (M j)"
-  shows "\<mu>' (emb I J (Pi\<^isub>E J X)) = (\<Prod>j\<in>J. M.\<mu>' j (X j))"
-proof cases
-  assume "J = {}"
-  then have "emb I J (Pi\<^isub>E J X) = space infprod_algebra"
-    by (auto simp: infprod_algebra_def generator_def sigma_def emb_def)
-  then show ?thesis using `J = {}` P.prob_space
-    by simp
-next
-  assume "J \<noteq> {}"
-  interpret J: finite_product_prob_space M J by default fact+
-  have "(\<Prod>i\<in>J. M.\<mu>' i (X i)) = J.\<mu>' (Pi\<^isub>E J X)"
-    using J `J \<noteq> {}` by (subst J.finite_measure_times) auto
-  also have "\<dots> = \<mu>' (emb I J (Pi\<^isub>E J X))"
-    using J `J \<noteq> {}` by (intro finite_measure_infprod_emb[symmetric]) auto
-  finally show ?thesis by simp
-qed
-
-lemma sigma_sets_mono: assumes "A \<subseteq> sigma_sets X B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
+sublocale product_prob_space \<subseteq> P: prob_space "Pi\<^isub>M I M"
 proof
-  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
-    by induct (insert `A \<subseteq> sigma_sets X B`, auto intro: sigma_sets.intros)
-qed
-
-lemma sigma_sets_mono': assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
-proof
-  fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
-    by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros)
+  show "emeasure (Pi\<^isub>M I M) (space (Pi\<^isub>M I M)) = 1"
+  proof cases
+    assume "I = {}" then show ?thesis by (simp add: space_PiM_empty)
+  next
+    assume "I \<noteq> {}"
+    then obtain i where "i \<in> I" by auto
+    moreover then have "emb I {i} (\<Pi>\<^isub>E i\<in>{i}. space (M i)) = (space (Pi\<^isub>M I M))"
+      by (auto simp: prod_emb_def space_PiM)
+    ultimately show ?thesis
+      using emeasure_PiM_emb_not_empty[of "{i}" "\<lambda>i. space (M i)"]
+      by (simp add: emeasure_PiM emeasure_space_1)
+  qed
 qed
 
-lemma sigma_sets_superset_generator: "A \<subseteq> sigma_sets X A"
-  by (auto intro: sigma_sets.Basic)
-
-lemma (in product_prob_space) infprod_algebra_alt:
-  "Pi\<^isub>P I M = sigma \<lparr> space = space (Pi\<^isub>P I M),
-    sets = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` Pi\<^isub>E J ` (\<Pi> i \<in> J. sets (M i))),
-    measure = measure (Pi\<^isub>P I M) \<rparr>"
-  (is "_ = sigma \<lparr> space = ?O, sets = ?M, measure = ?m \<rparr>")
-proof (rule measure_space.equality)
-  let ?G = "\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^isub>M J M)"
-  have "sigma_sets ?O ?M = sigma_sets ?O ?G"
-  proof (intro equalityI sigma_sets_mono UN_least)
-    fix J assume J: "J \<in> {J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}"
-    have "emb I J ` Pi\<^isub>E J ` (\<Pi> i\<in>J. sets (M i)) \<subseteq> emb I J ` sets (Pi\<^isub>M J M)" by auto
-    also have "\<dots> \<subseteq> ?G" using J by (rule UN_upper)
-    also have "\<dots> \<subseteq> sigma_sets ?O ?G" by (rule sigma_sets_superset_generator)
-    finally show "emb I J ` Pi\<^isub>E J ` (\<Pi> i\<in>J. sets (M i)) \<subseteq> sigma_sets ?O ?G" .
-    have "emb I J ` sets (Pi\<^isub>M J M) = emb I J ` sigma_sets (space (Pi\<^isub>M J M)) (Pi\<^isub>E J ` (\<Pi> i \<in> J. sets (M i)))"
-      by (simp add: sets_sigma product_algebra_generator_def product_algebra_def)
-    also have "\<dots> = sigma_sets (space (Pi\<^isub>M I M)) (emb I J ` Pi\<^isub>E J ` (\<Pi> i \<in> J. sets (M i)))"
-      using J M.sets_into_space
-      by (auto simp: emb_def [abs_def] intro!: sigma_sets_vimage[symmetric]) blast
-    also have "\<dots> \<subseteq> sigma_sets (space (Pi\<^isub>M I M)) ?M"
-      using J by (intro sigma_sets_mono') auto
-    finally show "emb I J ` sets (Pi\<^isub>M J M) \<subseteq> sigma_sets ?O ?M"
-      by (simp add: infprod_algebra_def generator_def)
-  qed
-  then show "sets (Pi\<^isub>P I M) = sets (sigma \<lparr> space = ?O, sets = ?M, measure = ?m \<rparr>)"
-    by (simp_all add: infprod_algebra_def generator_def sets_sigma)
-qed simp_all
-
-lemma (in product_prob_space) infprod_algebra_alt2:
-  "Pi\<^isub>P I M = sigma \<lparr> space = space (Pi\<^isub>P I M),
-    sets = (\<Union>i\<in>I. (\<lambda>A. (\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M)) ` sets (M i)),
-    measure = measure (Pi\<^isub>P I M) \<rparr>"
-  (is "_ = ?S")
-proof (rule measure_space.equality)
-  let "sigma \<lparr> space = ?O, sets = ?A, \<dots> = _ \<rparr>" = ?S
-  let ?G = "(\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` Pi\<^isub>E J ` (\<Pi> i \<in> J. sets (M i)))"
-  have "sets (Pi\<^isub>P I M) = sigma_sets ?O ?G"
-    by (subst infprod_algebra_alt) (simp add: sets_sigma)
-  also have "\<dots> = sigma_sets ?O ?A"
-  proof (intro equalityI sigma_sets_mono subsetI)
-    interpret A: sigma_algebra ?S
-      by (rule sigma_algebra_sigma) auto
-    fix A assume "A \<in> ?G"
-    then obtain J B where "finite J" "J \<noteq> {}" "J \<subseteq> I" "A = emb I J (Pi\<^isub>E J B)"
-        and B: "\<And>i. i \<in> J \<Longrightarrow> B i \<in> sets (M i)"
-      by auto
-    then have A: "A = (\<Inter>j\<in>J. (\<lambda>x. x j) -` (B j) \<inter> space (Pi\<^isub>P I M))"
-      by (auto simp: emb_def infprod_algebra_def generator_def Pi_iff)
-    { fix j assume "j\<in>J"
-      with `J \<subseteq> I` have "j \<in> I" by auto
-      with `j \<in> J` B have "(\<lambda>x. x j) -` (B j) \<inter> space (Pi\<^isub>P I M) \<in> sets ?S"
-        by (auto simp: sets_sigma intro: sigma_sets.Basic) }
-    with `finite J` `J \<noteq> {}` have "A \<in> sets ?S"
-      unfolding A by (intro A.finite_INT) auto
-    then show "A \<in> sigma_sets ?O ?A" by (simp add: sets_sigma)
-  next
-    fix A assume "A \<in> ?A"
-    then obtain i B where i: "i \<in> I" "B \<in> sets (M i)"
-        and "A = (\<lambda>x. x i) -` B \<inter> space (Pi\<^isub>P I M)"
-      by auto
-    then have "A = emb I {i} (Pi\<^isub>E {i} (\<lambda>_. B))"
-      by (auto simp: emb_def infprod_algebra_def generator_def Pi_iff)
-    with i show "A \<in> sigma_sets ?O ?G"
-      by (intro sigma_sets.Basic UN_I[where a="{i}"]) auto
-  qed
-  also have "\<dots> = sets ?S"
-    by (simp add: sets_sigma)
-  finally show "sets (Pi\<^isub>P I M) = sets ?S" .
-qed simp_all
-
-lemma (in product_prob_space) measurable_into_infprod_algebra:
-  assumes "sigma_algebra N"
-  assumes f: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. f x i) \<in> measurable N (M i)"
-  assumes ext: "\<And>x. x \<in> space N \<Longrightarrow> f x \<in> extensional I"
-  shows "f \<in> measurable N (Pi\<^isub>P I M)"
-proof -
-  interpret N: sigma_algebra N by fact
-  have f_in: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. f x i) \<in> space N \<rightarrow> space (M i)"
-    using f by (auto simp: measurable_def)
-  { fix i A assume i: "i \<in> I" "A \<in> sets (M i)"
-    then have "f -` (\<lambda>x. x i) -` A \<inter> f -` space infprod_algebra \<inter> space N = (\<lambda>x. f x i) -` A \<inter> space N"
-      using f_in ext by (auto simp: infprod_algebra_def generator_def)
-    also have "\<dots> \<in> sets N"
-      by (rule measurable_sets f i)+
-    finally have "f -` (\<lambda>x. x i) -` A \<inter> f -` space infprod_algebra \<inter> space N \<in> sets N" . }
-  with f_in ext show ?thesis
-    by (subst infprod_algebra_alt2)
-       (auto intro!: N.measurable_sigma simp: Pi_iff infprod_algebra_def generator_def)
+lemma (in product_prob_space) emeasure_PiM_emb:
+  assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
+  shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))"
+proof cases
+  assume "J = {}"
+  moreover have "emb I {} {\<lambda>x. undefined} = space (Pi\<^isub>M I M)"
+    by (auto simp: space_PiM prod_emb_def)
+  ultimately show ?thesis
+    by (simp add: space_PiM_empty P.emeasure_space_1)
+next
+  assume "J \<noteq> {}" with X show ?thesis
+    by (subst emeasure_PiM_emb_not_empty) (auto simp: emeasure_PiM)
 qed
 
-lemma (in product_prob_space) measurable_singleton_infprod:
-  assumes "i \<in> I"
-  shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>P I M) (M i)"
-proof (unfold measurable_def, intro CollectI conjI ballI)
-  show "(\<lambda>x. x i) \<in> space (Pi\<^isub>P I M) \<rightarrow> space (M i)"
-    using M.sets_into_space `i \<in> I`
-    by (auto simp: infprod_algebra_def generator_def)
-  fix A assume "A \<in> sets (M i)"
-  have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M) = emb I {i} (\<Pi>\<^isub>E _\<in>{i}. A)"
-    by (auto simp: infprod_algebra_def generator_def emb_def)
-  also have "\<dots> \<in> sets (Pi\<^isub>P I M)"
-    using `i \<in> I` `A \<in> sets (M i)`
-    by (intro emb_in_infprod_algebra product_algebraI) auto
-  finally show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M) \<in> sets (Pi\<^isub>P I M)" .
-qed
+lemma (in product_prob_space) measure_PiM_emb:
+  assumes "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
+  shows "measure (PiM I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))"
+  using emeasure_PiM_emb[OF assms]
+  unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal)
 
-lemma (in product_prob_space) sigma_product_algebra_sigma_eq:
-  assumes M: "\<And>i. i \<in> I \<Longrightarrow> M i = sigma (E i)"
-  shows "sets (Pi\<^isub>P I M) = sigma_sets (space (Pi\<^isub>P I M)) (\<Union>i\<in>I. (\<lambda>A. (\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M)) ` sets (E i))"
-proof -
-  let ?E = "(\<Union>i\<in>I. (\<lambda>A. (\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M)) ` sets (E i))"
-  let ?M = "(\<Union>i\<in>I. (\<lambda>A. (\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M)) ` sets (M i))"
-  { fix i A assume "i\<in>I" "A \<in> sets (E i)"
-    then have "A \<in> sets (M i)" using M by auto
-    then have "A \<in> Pow (space (M i))" using M.sets_into_space by auto
-    then have "A \<in> Pow (space (E i))" using M[OF `i \<in> I`] by auto }
-  moreover
-  have "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. x i) \<in> space infprod_algebra \<rightarrow> space (E i)"
-    by (auto simp: M infprod_algebra_def generator_def Pi_iff)
-  ultimately have "sigma_sets (space (Pi\<^isub>P I M)) ?M \<subseteq> sigma_sets (space (Pi\<^isub>P I M)) ?E"
-    apply (intro sigma_sets_mono UN_least)
-    apply (simp add: sets_sigma M)
-    apply (subst sigma_sets_vimage[symmetric])
-    apply (auto intro!: sigma_sets_mono')
-    done
-  moreover have "sigma_sets (space (Pi\<^isub>P I M)) ?E \<subseteq> sigma_sets (space (Pi\<^isub>P I M)) ?M"
-    by (intro sigma_sets_mono') (auto simp: M)
-  ultimately show ?thesis
-    by (subst infprod_algebra_alt2) (auto simp: sets_sigma)
-qed
-
-lemma (in product_prob_space) Int_proj_eq_emb:
-  assumes "J \<noteq> {}" "J \<subseteq> I"
-  shows "(\<Inter>i\<in>J. (\<lambda>x. x i) -` A i \<inter> space (Pi\<^isub>P I M)) = emb I J (Pi\<^isub>E J A)"
-  using assms by (auto simp: infprod_algebra_def generator_def emb_def Pi_iff)
-
-lemma (in product_prob_space) emb_insert:
-  "i \<notin> J \<Longrightarrow> emb I J (Pi\<^isub>E J f) \<inter> ((\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>P I M)) =
-    emb I (insert i J) (Pi\<^isub>E (insert i J) (f(i := A)))"
-  by (auto simp: emb_def Pi_iff infprod_algebra_def generator_def split: split_if_asm)
+lemma (in finite_product_prob_space) finite_measure_PiM_emb:
+  "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> measure (PiM I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))"
+  using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets_into_space, of I A M]
+  by auto
 
 subsection {* Sequence space *}
 
@@ -1003,36 +722,30 @@
 
 lemma (in sequence_space) infprod_in_sets[intro]:
   fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets (M i)"
-  shows "Pi UNIV E \<in> sets (Pi\<^isub>P UNIV M)"
+  shows "Pi UNIV E \<in> sets (Pi\<^isub>M UNIV M)"
 proof -
   have "Pi UNIV E = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^isub>E j\<in>{..i}. E j))"
-    using E E[THEN M.sets_into_space]
-    by (auto simp: emb_def Pi_iff extensional_def) blast
-  with E show ?thesis
-    by (auto intro: emb_in_infprod_algebra)
+    using E E[THEN sets_into_space]
+    by (auto simp: prod_emb_def Pi_iff extensional_def) blast
+  with E show ?thesis by auto
 qed
 
-lemma (in sequence_space) measure_infprod:
+lemma (in sequence_space) measure_PiM_countable:
   fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets (M i)"
-  shows "(\<lambda>n. \<Prod>i\<le>n. M.\<mu>' i (E i)) ----> \<mu>' (Pi UNIV E)"
+  shows "(\<lambda>n. \<Prod>i\<le>n. measure (M i) (E i)) ----> measure (Pi\<^isub>M UNIV M) (Pi UNIV E)"
 proof -
   let ?E = "\<lambda>n. emb UNIV {..n} (Pi\<^isub>E {.. n} E)"
-  { fix n :: nat
-    interpret n: finite_product_prob_space M "{..n}" by default auto
-    have "(\<Prod>i\<le>n. M.\<mu>' i (E i)) = n.\<mu>' (Pi\<^isub>E {.. n} E)"
-      using E by (subst n.finite_measure_times) auto
-    also have "\<dots> = \<mu>' (?E n)"
-      using E by (intro finite_measure_infprod_emb[symmetric]) auto
-    finally have "(\<Prod>i\<le>n. M.\<mu>' i (E i)) = \<mu>' (?E n)" . }
+  have "\<And>n. (\<Prod>i\<le>n. measure (M i) (E i)) = measure (Pi\<^isub>M UNIV M) (?E n)"
+    using E by (simp add: measure_PiM_emb)
   moreover have "Pi UNIV E = (\<Inter>n. ?E n)"
-    using E E[THEN M.sets_into_space]
-    by (auto simp: emb_def extensional_def Pi_iff) blast
-  moreover have "range ?E \<subseteq> sets (Pi\<^isub>P UNIV M)"
+    using E E[THEN sets_into_space]
+    by (auto simp: prod_emb_def extensional_def Pi_iff) blast
+  moreover have "range ?E \<subseteq> sets (Pi\<^isub>M UNIV M)"
     using E by auto
   moreover have "decseq ?E"
-    by (auto simp: emb_def Pi_iff decseq_def)
+    by (auto simp: prod_emb_def Pi_iff decseq_def)
   ultimately show ?thesis
-    by (simp add: finite_continuity_from_above)
+    by (simp add: finite_Lim_measure_decseq)
 qed
 
 end
\ No newline at end of file
--- a/src/HOL/Probability/Information.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Information.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -172,147 +172,152 @@
 Kullback$-$Leibler distance. *}
 
 definition
-  "entropy_density b M \<nu> = log b \<circ> real \<circ> RN_deriv M \<nu>"
+  "entropy_density b M N = log b \<circ> real \<circ> RN_deriv M N"
 
 definition
-  "KL_divergence b M \<nu> = integral\<^isup>L (M\<lparr>measure := \<nu>\<rparr>) (entropy_density b M \<nu>)"
+  "KL_divergence b M N = integral\<^isup>L N (entropy_density b M N)"
 
 lemma (in information_space) measurable_entropy_density:
-  assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)"
-  assumes ac: "absolutely_continuous \<nu>"
-  shows "entropy_density b M \<nu> \<in> borel_measurable M"
+  assumes ac: "absolutely_continuous M N" "sets N = events"
+  shows "entropy_density b M N \<in> borel_measurable M"
 proof -
-  interpret \<nu>: prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact
-  have "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by fact
-  from RN_deriv[OF this ac] b_gt_1 show ?thesis
+  from borel_measurable_RN_deriv[OF ac] b_gt_1 show ?thesis
     unfolding entropy_density_def
     by (intro measurable_comp) auto
 qed
 
-lemma (in information_space) KL_gt_0:
-  assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)"
-  assumes ac: "absolutely_continuous \<nu>"
-  assumes int: "integrable (M\<lparr> measure := \<nu> \<rparr>) (entropy_density b M \<nu>)"
-  assumes A: "A \<in> sets M" "\<nu> A \<noteq> \<mu> A"
-  shows "0 < KL_divergence b M \<nu>"
-proof -
-  interpret \<nu>: prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact
-  have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
-  have fms: "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" by unfold_locales
-  note RN = RN_deriv[OF ms ac]
-
-  from real_RN_deriv[OF fms ac] guess D . note D = this
-  with absolutely_continuous_AE[OF ms] ac
-  have D\<nu>: "AE x in M\<lparr>measure := \<nu>\<rparr>. RN_deriv M \<nu> x = ereal (D x)"
-    by auto
-
-  def f \<equiv> "\<lambda>x. if D x = 0 then 1 else 1 / D x"
-  with D have f_borel: "f \<in> borel_measurable M"
-    by (auto intro!: measurable_If)
+lemma (in sigma_finite_measure) KL_density:
+  fixes f :: "'a \<Rightarrow> real"
+  assumes "1 < b"
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  shows "KL_divergence b M (density M f) = (\<integral>x. f x * log b (f x) \<partial>M)"
+  unfolding KL_divergence_def
+proof (subst integral_density)
+  show "entropy_density b M (density M (\<lambda>x. ereal (f x))) \<in> borel_measurable M"
+    using f `1 < b`
+    by (auto simp: comp_def entropy_density_def intro!: borel_measurable_log borel_measurable_RN_deriv_density)
+  have "density M (RN_deriv M (density M f)) = density M f"
+    using f by (intro density_RN_deriv_density) auto
+  then have eq: "AE x in M. RN_deriv M (density M f) x = f x"
+    using f
+    by (intro density_unique)
+       (auto intro!: borel_measurable_log borel_measurable_RN_deriv_density simp: RN_deriv_density_nonneg)
+  show "(\<integral>x. f x * entropy_density b M (density M (\<lambda>x. ereal (f x))) x \<partial>M) = (\<integral>x. f x * log b (f x) \<partial>M)"
+    apply (intro integral_cong_AE)
+    using eq
+    apply eventually_elim
+    apply (auto simp: entropy_density_def)
+    done
+qed fact+
 
-  have "KL_divergence b M \<nu> = 1 / ln b * (\<integral> x. ln b * entropy_density b M \<nu> x \<partial>M\<lparr>measure := \<nu>\<rparr>)"
-    unfolding KL_divergence_def using int b_gt_1
-    by (simp add: integral_cmult)
-
-  { fix A assume "A \<in> sets M"
-    with RN D have "\<nu>.\<mu> A = (\<integral>\<^isup>+ x. ereal (D x) * indicator A x \<partial>M)"
-      by (auto intro!: positive_integral_cong_AE) }
-  note D_density = this
+lemma (in sigma_finite_measure) KL_density_density:
+  fixes f g :: "'a \<Rightarrow> real"
+  assumes "1 < b"
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
+  assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0"
+  shows "KL_divergence b (density M f) (density M g) = (\<integral>x. g x * log b (g x / f x) \<partial>M)"
+proof -
+  interpret Mf: sigma_finite_measure "density M f"
+    using f by (subst sigma_finite_iff_density_finite) auto
+  have "KL_divergence b (density M f) (density M g) =
+    KL_divergence b (density M f) (density (density M f) (\<lambda>x. g x / f x))"
+    using f g ac by (subst density_density_divide) simp_all
+  also have "\<dots> = (\<integral>x. (g x / f x) * log b (g x / f x) \<partial>density M f)"
+    using f g `1 < b` by (intro Mf.KL_density) (auto simp: AE_density divide_nonneg_nonneg)
+  also have "\<dots> = (\<integral>x. g x * log b (g x / f x) \<partial>M)"
+    using ac f g `1 < b` by (subst integral_density) (auto intro!: integral_cong_AE)
+  finally show ?thesis .
+qed
 
-  have ln_entropy: "(\<lambda>x. ln b * entropy_density b M \<nu> x) \<in> borel_measurable M"
-    using measurable_entropy_density[OF ps ac] by auto
+lemma (in information_space) KL_gt_0:
+  fixes D :: "'a \<Rightarrow> real"
+  assumes "prob_space (density M D)"
+  assumes D: "D \<in> borel_measurable M" "AE x in M. 0 \<le> D x"
+  assumes int: "integrable M (\<lambda>x. D x * log b (D x))"
+  assumes A: "density M D \<noteq> M"
+  shows "0 < KL_divergence b M (density M D)"
+proof -
+  interpret N: prob_space "density M D" by fact
 
-  have "integrable (M\<lparr>measure := \<nu>\<rparr>) (\<lambda>x. ln b * entropy_density b M \<nu> x)"
-    using int by auto
-  moreover have "integrable (M\<lparr>measure := \<nu>\<rparr>) (\<lambda>x. ln b * entropy_density b M \<nu> x) \<longleftrightarrow>
-      integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))"
-    using D D_density ln_entropy
-    by (intro integral_translated_density) auto
-  ultimately have M_int: "integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))"
-    by simp
+  obtain A where "A \<in> sets M" "emeasure (density M D) A \<noteq> emeasure M A"
+    using measure_eqI[of "density M D" M] `density M D \<noteq> M` by auto
+
+  let ?D_set = "{x\<in>space M. D x \<noteq> 0}"
+  have [simp, intro]: "?D_set \<in> sets M"
+    using D by auto
 
   have D_neg: "(\<integral>\<^isup>+ x. ereal (- D x) \<partial>M) = 0"
     using D by (subst positive_integral_0_iff_AE) auto
 
-  have "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = \<nu> (space M)"
-    using RN D by (auto intro!: positive_integral_cong_AE)
+  have "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = emeasure (density M D) (space M)"
+    using D by (simp add: emeasure_density cong: positive_integral_cong)
   then have D_pos: "(\<integral>\<^isup>+ x. ereal (D x) \<partial>M) = 1"
-    using \<nu>.measure_space_1 by simp
-
-  have "integrable M D"
-    using D_pos D_neg D by (auto simp: integrable_def)
+    using N.emeasure_space_1 by simp
 
-  have "integral\<^isup>L M D = 1"
-    using D_pos D_neg by (auto simp: lebesgue_integral_def)
+  have "integrable M D" "integral\<^isup>L M D = 1"
+    using D D_pos D_neg unfolding integrable_def lebesgue_integral_def by simp_all
 
-  let ?D_set = "{x\<in>space M. D x \<noteq> 0}"
-  have [simp, intro]: "?D_set \<in> sets M"
-    using D by (auto intro: sets_Collect)
-
-  have "0 \<le> 1 - \<mu>' ?D_set"
+  have "0 \<le> 1 - measure M ?D_set"
     using prob_le_1 by (auto simp: field_simps)
   also have "\<dots> = (\<integral> x. D x - indicator ?D_set x \<partial>M)"
     using `integrable M D` `integral\<^isup>L M D = 1`
-    by (simp add: \<mu>'_def)
-  also have "\<dots> < (\<integral> x. D x * (ln b * entropy_density b M \<nu> x) \<partial>M)"
+    by (simp add: emeasure_eq_measure)
+  also have "\<dots> < (\<integral> x. D x * (ln b * log b (D x)) \<partial>M)"
   proof (rule integral_less_AE)
     show "integrable M (\<lambda>x. D x - indicator ?D_set x)"
       using `integrable M D`
       by (intro integral_diff integral_indicator) auto
   next
-    show "integrable M (\<lambda>x. D x * (ln b * entropy_density b M \<nu> x))"
-      by fact
+    from integral_cmult(1)[OF int, of "ln b"]
+    show "integrable M (\<lambda>x. D x * (ln b * log b (D x)))" 
+      by (simp add: ac_simps)
   next
-    show "\<mu> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<noteq> 0"
+    show "emeasure M {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<noteq> 0"
     proof
-      assume eq_0: "\<mu> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} = 0"
-      then have disj: "AE x. D x = 1 \<or> D x = 0"
+      assume eq_0: "emeasure M {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} = 0"
+      then have disj: "AE x in M. D x = 1 \<or> D x = 0"
         using D(1) by (auto intro!: AE_I[OF subset_refl] sets_Collect)
 
-      have "\<mu> {x\<in>space M. D x = 1} = (\<integral>\<^isup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)"
+      have "emeasure M {x\<in>space M. D x = 1} = (\<integral>\<^isup>+ x. indicator {x\<in>space M. D x = 1} x \<partial>M)"
         using D(1) by auto
-      also have "\<dots> = (\<integral>\<^isup>+ x. ereal (D x) * indicator {x\<in>space M. D x \<noteq> 0} x \<partial>M)"
+      also have "\<dots> = (\<integral>\<^isup>+ x. ereal (D x) \<partial>M)"
         using disj by (auto intro!: positive_integral_cong_AE simp: indicator_def one_ereal_def)
-      also have "\<dots> = \<nu> {x\<in>space M. D x \<noteq> 0}"
-        using D(1) D_density by auto
-      also have "\<dots> = \<nu> (space M)"
-        using D_density D(1) by (auto intro!: positive_integral_cong simp: indicator_def)
-      finally have "AE x. D x = 1"
-        using D(1) \<nu>.measure_space_1 by (intro AE_I_eq_1) auto
+      finally have "AE x in M. D x = 1"
+        using D D_pos by (intro AE_I_eq_1) auto
       then have "(\<integral>\<^isup>+x. indicator A x\<partial>M) = (\<integral>\<^isup>+x. ereal (D x) * indicator A x\<partial>M)"
         by (intro positive_integral_cong_AE) (auto simp: one_ereal_def[symmetric])
-      also have "\<dots> = \<nu> A"
-        using `A \<in> sets M` D_density by simp
-      finally show False using `A \<in> sets M` `\<nu> A \<noteq> \<mu> A` by simp
+      also have "\<dots> = density M D A"
+        using `A \<in> sets M` D by (simp add: emeasure_density)
+      finally show False using `A \<in> sets M` `emeasure (density M D) A \<noteq> emeasure M A` by simp
     qed
     show "{x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<in> sets M"
-      using D(1) by (auto intro: sets_Collect)
+      using D(1) by (auto intro: sets_Collect_conj)
 
-    show "AE t. t \<in> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<longrightarrow>
-      D t - indicator ?D_set t \<noteq> D t * (ln b * entropy_density b M \<nu> t)"
+    show "AE t in M. t \<in> {x\<in>space M. D x \<noteq> 1 \<and> D x \<noteq> 0} \<longrightarrow>
+      D t - indicator ?D_set t \<noteq> D t * (ln b * log b (D t))"
       using D(2)
-    proof (elim AE_mp, safe intro!: AE_I2)
-      fix t assume Dt: "t \<in> space M" "D t \<noteq> 1" "D t \<noteq> 0"
-        and RN: "RN_deriv M \<nu> t = ereal (D t)"
-        and eq: "D t - indicator ?D_set t = D t * (ln b * entropy_density b M \<nu> t)"
+    proof (eventually_elim, safe)
+      fix t assume Dt: "t \<in> space M" "D t \<noteq> 1" "D t \<noteq> 0" "0 \<le> D t"
+        and eq: "D t - indicator ?D_set t = D t * (ln b * log b (D t))"
 
       have "D t - 1 = D t - indicator ?D_set t"
         using Dt by simp
       also note eq
-      also have "D t * (ln b * entropy_density b M \<nu> t) = - D t * ln (1 / D t)"
-        using RN b_gt_1 `D t \<noteq> 0` `0 \<le> D t`
-        by (simp add: entropy_density_def log_def ln_div less_le)
+      also have "D t * (ln b * log b (D t)) = - D t * ln (1 / D t)"
+        using b_gt_1 `D t \<noteq> 0` `0 \<le> D t`
+        by (simp add: log_def ln_div less_le)
       finally have "ln (1 / D t) = 1 / D t - 1"
         using `D t \<noteq> 0` by (auto simp: field_simps)
       from ln_eq_minus_one[OF _ this] `D t \<noteq> 0` `0 \<le> D t` `D t \<noteq> 1`
       show False by auto
     qed
 
-    show "AE t. D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)"
-      using D(2)
-    proof (elim AE_mp, intro AE_I2 impI)
-      fix t assume "t \<in> space M" and RN: "RN_deriv M \<nu> t = ereal (D t)"
-      show "D t - indicator ?D_set t \<le> D t * (ln b * entropy_density b M \<nu> t)"
+    show "AE t in M. D t - indicator ?D_set t \<le> D t * (ln b * log b (D t))"
+      using D(2) AE_space
+    proof eventually_elim
+      fix t assume "t \<in> space M" "0 \<le> D t"
+      show "D t - indicator ?D_set t \<le> D t * (ln b * log b (D t))"
       proof cases
         assume asm: "D t \<noteq> 0"
         then have "0 < D t" using `0 \<le> D t` by auto
@@ -321,592 +326,425 @@
           using asm `t \<in> space M` by (simp add: field_simps)
         also have "- D t * (1 / D t - 1) \<le> - D t * ln (1 / D t)"
           using ln_le_minus_one `0 < 1 / D t` by (intro mult_left_mono_neg) auto
-        also have "\<dots> = D t * (ln b * entropy_density b M \<nu> t)"
-          using `0 < D t` RN b_gt_1
-          by (simp_all add: log_def ln_div entropy_density_def)
+        also have "\<dots> = D t * (ln b * log b (D t))"
+          using `0 < D t` b_gt_1
+          by (simp_all add: log_def ln_div)
         finally show ?thesis by simp
       qed simp
     qed
   qed
-  also have "\<dots> = (\<integral> x. ln b * entropy_density b M \<nu> x \<partial>M\<lparr>measure := \<nu>\<rparr>)"
-    using D D_density ln_entropy
-    by (intro integral_translated_density[symmetric]) auto
-  also have "\<dots> = ln b * (\<integral> x. entropy_density b M \<nu> x \<partial>M\<lparr>measure := \<nu>\<rparr>)"
-    using int by (rule \<nu>.integral_cmult)
-  finally show "0 < KL_divergence b M \<nu>"
-    using b_gt_1 by (auto simp: KL_divergence_def zero_less_mult_iff)
+  also have "\<dots> = (\<integral> x. ln b * (D x * log b (D x)) \<partial>M)"
+    by (simp add: ac_simps)
+  also have "\<dots> = ln b * (\<integral> x. D x * log b (D x) \<partial>M)"
+    using int by (rule integral_cmult)
+  finally show ?thesis
+    using b_gt_1 D by (subst KL_density) (auto simp: zero_less_mult_iff)
 qed
 
-lemma (in sigma_finite_measure) KL_eq_0:
-  assumes eq: "\<forall>A\<in>sets M. \<nu> A = measure M A"
-  shows "KL_divergence b M \<nu> = 0"
+lemma (in sigma_finite_measure) KL_same_eq_0: "KL_divergence b M M = 0"
 proof -
-  have "AE x. 1 = RN_deriv M \<nu> x"
+  have "AE x in M. 1 = RN_deriv M M x"
   proof (rule RN_deriv_unique)
-    show "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
-      using eq by (intro measure_space_cong) auto
-    show "absolutely_continuous \<nu>"
-      unfolding absolutely_continuous_def using eq by auto
-    show "(\<lambda>x. 1) \<in> borel_measurable M" "AE x. 0 \<le> (1 :: ereal)" by auto
-    fix A assume "A \<in> sets M"
-    with eq show "\<nu> A = \<integral>\<^isup>+ x. 1 * indicator A x \<partial>M" by simp
+    show "(\<lambda>x. 1) \<in> borel_measurable M" "AE x in M. 0 \<le> (1 :: ereal)" by auto
+    show "density M (\<lambda>x. 1) = M"
+      apply (auto intro!: measure_eqI emeasure_density)
+      apply (subst emeasure_density)
+      apply auto
+      done
   qed
-  then have "AE x. log b (real (RN_deriv M \<nu> x)) = 0"
+  then have "AE x in M. log b (real (RN_deriv M M x)) = 0"
     by (elim AE_mp) simp
   from integral_cong_AE[OF this]
-  have "integral\<^isup>L M (entropy_density b M \<nu>) = 0"
+  have "integral\<^isup>L M (entropy_density b M M) = 0"
     by (simp add: entropy_density_def comp_def)
-  with eq show "KL_divergence b M \<nu> = 0"
+  then show "KL_divergence b M M = 0"
     unfolding KL_divergence_def
-    by (subst integral_cong_measure) auto
+    by auto
 qed
 
-lemma (in information_space) KL_eq_0_imp:
-  assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)"
-  assumes ac: "absolutely_continuous \<nu>"
-  assumes int: "integrable (M\<lparr> measure := \<nu> \<rparr>) (entropy_density b M \<nu>)"
-  assumes KL: "KL_divergence b M \<nu> = 0"
-  shows "\<forall>A\<in>sets M. \<nu> A = \<mu> A"
-  by (metis less_imp_neq KL_gt_0 assms)
+lemma (in information_space) KL_eq_0_iff_eq:
+  fixes D :: "'a \<Rightarrow> real"
+  assumes "prob_space (density M D)"
+  assumes D: "D \<in> borel_measurable M" "AE x in M. 0 \<le> D x"
+  assumes int: "integrable M (\<lambda>x. D x * log b (D x))"
+  shows "KL_divergence b M (density M D) = 0 \<longleftrightarrow> density M D = M"
+  using KL_same_eq_0[of b] KL_gt_0[OF assms]
+  by (auto simp: less_le)
 
-lemma (in information_space) KL_ge_0:
-  assumes ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)"
-  assumes ac: "absolutely_continuous \<nu>"
-  assumes int: "integrable (M\<lparr> measure := \<nu> \<rparr>) (entropy_density b M \<nu>)"
-  shows "0 \<le> KL_divergence b M \<nu>"
-  using KL_eq_0 KL_gt_0[OF ps ac int]
-  by (cases "\<forall>A\<in>sets M. \<nu> A = measure M A") (auto simp: le_less)
-
-
-lemma (in sigma_finite_measure) KL_divergence_vimage:
-  assumes T: "T \<in> measure_preserving M M'"
-    and T': "T' \<in> measure_preserving (M'\<lparr> measure := \<nu>' \<rparr>) (M\<lparr> measure := \<nu> \<rparr>)"
-    and inv: "\<And>x. x \<in> space M \<Longrightarrow> T' (T x) = x"
-    and inv': "\<And>x. x \<in> space M' \<Longrightarrow> T (T' x) = x"
-  and \<nu>': "measure_space (M'\<lparr>measure := \<nu>'\<rparr>)" "measure_space.absolutely_continuous M' \<nu>'"
-  and "1 < b"
-  shows "KL_divergence b M' \<nu>' = KL_divergence b M \<nu>"
+lemma (in information_space) KL_eq_0_iff_eq_ac:
+  fixes D :: "'a \<Rightarrow> real"
+  assumes "prob_space N"
+  assumes ac: "absolutely_continuous M N" "sets N = sets M"
+  assumes int: "integrable N (entropy_density b M N)"
+  shows "KL_divergence b M N = 0 \<longleftrightarrow> N = M"
 proof -
-  interpret \<nu>': measure_space "M'\<lparr>measure := \<nu>'\<rparr>" by fact
-  have M: "measure_space (M\<lparr> measure := \<nu>\<rparr>)"
-    by (rule \<nu>'.measure_space_vimage[OF _ T'], simp) default
-  have "sigma_algebra (M'\<lparr> measure := \<nu>'\<rparr>)" by default
-  then have saM': "sigma_algebra M'" by simp
-  then interpret M': measure_space M' by (rule measure_space_vimage) fact
-  have ac: "absolutely_continuous \<nu>" unfolding absolutely_continuous_def
-  proof safe
-    fix N assume N: "N \<in> sets M" and N_0: "\<mu> N = 0"
-    then have N': "T' -` N \<inter> space M' \<in> sets M'"
-      using T' by (auto simp: measurable_def measure_preserving_def)
-    have "T -` (T' -` N \<inter> space M') \<inter> space M = N"
-      using inv T N sets_into_space[OF N] by (auto simp: measurable_def measure_preserving_def)
-    then have "measure M' (T' -` N \<inter> space M') = 0"
-      using measure_preservingD[OF T N'] N_0 by auto
-    with \<nu>'(2) N' show "\<nu> N = 0" using measure_preservingD[OF T', of N] N
-      unfolding M'.absolutely_continuous_def measurable_def by auto
-  qed
+  interpret N: prob_space N by fact
+  have "finite_measure N" by unfold_locales
+  from real_RN_deriv[OF this ac] guess D . note D = this
+  
+  have "N = density M (RN_deriv M N)"
+    using ac by (rule density_RN_deriv[symmetric])
+  also have "\<dots> = density M D"
+    using borel_measurable_RN_deriv[OF ac] D by (auto intro!: density_cong)
+  finally have N: "N = density M D" .
 
-  have sa: "sigma_algebra (M\<lparr>measure := \<nu>\<rparr>)" by simp default
-  have AE: "AE x. RN_deriv M' \<nu>' (T x) = RN_deriv M \<nu> x"
-    by (rule RN_deriv_vimage[OF T T' inv \<nu>'])
-  show ?thesis
-    unfolding KL_divergence_def entropy_density_def comp_def
-  proof (subst \<nu>'.integral_vimage[OF sa T'])
-    show "(\<lambda>x. log b (real (RN_deriv M \<nu> x))) \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)"
-      by (auto intro!: RN_deriv[OF M ac] borel_measurable_log[OF _ `1 < b`])
-    have "(\<integral> x. log b (real (RN_deriv M' \<nu>' x)) \<partial>M'\<lparr>measure := \<nu>'\<rparr>) =
-      (\<integral> x. log b (real (RN_deriv M' \<nu>' (T (T' x)))) \<partial>M'\<lparr>measure := \<nu>'\<rparr>)" (is "?l = _")
-      using inv' by (auto intro!: \<nu>'.integral_cong)
-    also have "\<dots> = (\<integral> x. log b (real (RN_deriv M \<nu> (T' x))) \<partial>M'\<lparr>measure := \<nu>'\<rparr>)" (is "_ = ?r")
-      using M ac AE
-      by (intro \<nu>'.integral_cong_AE \<nu>'.almost_everywhere_vimage[OF sa T'] absolutely_continuous_AE[OF M])
-         (auto elim!: AE_mp)
-    finally show "?l = ?r" .
-  qed
+  from absolutely_continuous_AE[OF ac(2,1) D(2)] D b_gt_1 ac measurable_entropy_density
+  have "integrable N (\<lambda>x. log b (D x))"
+    by (intro integrable_cong_AE[THEN iffD2, OF _ _ _ int])
+       (auto simp: N entropy_density_def)
+  with D b_gt_1 have "integrable M (\<lambda>x. D x * log b (D x))"
+    by (subst integral_density(2)[symmetric]) (auto simp: N[symmetric] comp_def)
+  with `prob_space N` D show ?thesis
+    unfolding N
+    by (intro KL_eq_0_iff_eq) auto
 qed
 
-lemma (in sigma_finite_measure) KL_divergence_cong:
-  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?\<nu>")
-  assumes [simp]: "sets N = sets M" "space N = space M"
-    "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A"
-    "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<nu>' A"
-  shows "KL_divergence b M \<nu> = KL_divergence b N \<nu>'"
-proof -
-  interpret \<nu>: measure_space ?\<nu> by fact
-  have "KL_divergence b M \<nu> = \<integral>x. log b (real (RN_deriv N \<nu>' x)) \<partial>?\<nu>"
-    by (simp cong: RN_deriv_cong \<nu>.integral_cong add: KL_divergence_def entropy_density_def)
-  also have "\<dots> = KL_divergence b N \<nu>'"
-    by (auto intro!: \<nu>.integral_cong_measure[symmetric] simp: KL_divergence_def entropy_density_def comp_def)
-  finally show ?thesis .
-qed
+lemma (in information_space) KL_nonneg:
+  assumes "prob_space (density M D)"
+  assumes D: "D \<in> borel_measurable M" "AE x in M. 0 \<le> D x"
+  assumes int: "integrable M (\<lambda>x. D x * log b (D x))"
+  shows "0 \<le> KL_divergence b M (density M D)"
+  using KL_gt_0[OF assms] by (cases "density M D = M") (auto simp: KL_same_eq_0)
 
-lemma (in finite_measure_space) KL_divergence_eq_finite:
-  assumes v: "finite_measure_space (M\<lparr>measure := \<nu>\<rparr>)"
-  assumes ac: "absolutely_continuous \<nu>"
-  shows "KL_divergence b M \<nu> = (\<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] entropy_density_def)
-  interpret v: finite_measure_space "M\<lparr>measure := \<nu>\<rparr>" by fact
-  have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
-  show "(\<Sum>x \<in> space M. log b (real (RN_deriv M \<nu> x)) * real (\<nu> {x})) = ?sum"
-    using RN_deriv_finite_measure[OF ms ac]
-    by (auto intro!: setsum_cong simp: field_simps)
-qed
+lemma (in sigma_finite_measure) KL_density_density_nonneg:
+  fixes f g :: "'a \<Rightarrow> real"
+  assumes "1 < b"
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "prob_space (density M f)"
+  assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" "prob_space (density M g)"
+  assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0"
+  assumes int: "integrable M (\<lambda>x. g x * log b (g x / f x))"
+  shows "0 \<le> KL_divergence b (density M f) (density M g)"
+proof -
+  interpret Mf: prob_space "density M f" by fact
+  interpret Mf: information_space "density M f" b by default fact
+  have eq: "density (density M f) (\<lambda>x. g x / f x) = density M g" (is "?DD = _")
+    using f g ac by (subst density_density_divide) simp_all
 
-lemma (in finite_prob_space) KL_divergence_positive_finite:
-  assumes v: "finite_prob_space (M\<lparr>measure := \<nu>\<rparr>)"
-  assumes ac: "absolutely_continuous \<nu>"
-  and "1 < b"
-  shows "0 \<le> KL_divergence b M \<nu>"
-proof -
-  interpret information_space M by default fact
-  interpret v: finite_prob_space "M\<lparr>measure := \<nu>\<rparr>" by fact
-  have ps: "prob_space (M\<lparr>measure := \<nu>\<rparr>)" by unfold_locales
-  from KL_ge_0[OF this ac v.integral_finite_singleton(1)] show ?thesis .
+  have "0 \<le> KL_divergence b (density M f) (density (density M f) (\<lambda>x. g x / f x))"
+  proof (rule Mf.KL_nonneg)
+    show "prob_space ?DD" unfolding eq by fact
+    from f g show "(\<lambda>x. g x / f x) \<in> borel_measurable (density M f)"
+      by auto
+    show "AE x in density M f. 0 \<le> g x / f x"
+      using f g by (auto simp: AE_density divide_nonneg_nonneg)
+    show "integrable (density M f) (\<lambda>x. g x / f x * log b (g x / f x))"
+      using `1 < b` f g ac
+      by (subst integral_density)
+         (auto intro!: integrable_cong_AE[THEN iffD2, OF _ _ _ int] measurable_If)
+  qed
+  also have "\<dots> = KL_divergence b (density M f) (density M g)"
+    using f g ac by (subst density_density_divide) simp_all
+  finally show ?thesis .
 qed
 
 subsection {* Mutual Information *}
 
 definition (in prob_space)
   "mutual_information b S T X Y =
-    KL_divergence b (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>)
-      (ereal\<circ>joint_distribution X Y)"
+    KL_divergence b (distr M S X \<Otimes>\<^isub>M distr M T Y) (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)))"
 
-lemma (in information_space)
+lemma (in information_space) mutual_information_indep_vars:
   fixes S T X Y
-  defines "P \<equiv> S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
+  defines "P \<equiv> distr M S X \<Otimes>\<^isub>M distr M T Y"
+  defines "Q \<equiv> distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
   shows "indep_var S X T Y \<longleftrightarrow>
     (random_variable S X \<and> random_variable T Y \<and>
-      measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y) \<and>
-      integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>)
-        (entropy_density b P (ereal\<circ>joint_distribution X Y)) \<and>
-     mutual_information b S T X Y = 0)"
+      absolutely_continuous P Q \<and> integrable Q (entropy_density b P Q) \<and>
+      mutual_information b S T X Y = 0)"
+  unfolding indep_var_distribution_eq
 proof safe
-  assume indep: "indep_var S X T Y"
-  then have "random_variable S X" "random_variable T Y"
-    by (blast dest: indep_var_rv1 indep_var_rv2)+
-  then show "sigma_algebra S" "X \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T"
-    by blast+
-
-  interpret X: prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
-    by (rule distribution_prob_space) fact
-  interpret Y: prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
-    by (rule distribution_prob_space) fact
-  interpret XY: pair_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default
-  interpret XY: information_space XY.P b by default (rule b_gt_1)
-
-  let ?J = "XY.P\<lparr> measure := (ereal\<circ>joint_distribution X Y) \<rparr>"
-  { fix A assume "A \<in> sets XY.P"
-    then have "ereal (joint_distribution X Y A) = XY.\<mu> A"
-      using indep_var_distributionD[OF indep]
-      by (simp add: XY.P.finite_measure_eq) }
-  note j_eq = this
-
-  interpret J: prob_space ?J
-    using j_eq by (intro XY.prob_space_cong) auto
-
-  have ac: "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)"
-    by (simp add: XY.absolutely_continuous_def j_eq)
-  then show "measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y)"
-    unfolding P_def .
-
-  have ed: "entropy_density b XY.P (ereal\<circ>joint_distribution X Y) \<in> borel_measurable XY.P"
-    by (rule XY.measurable_entropy_density) (default | fact)+
+  assume rv: "random_variable S X" "random_variable T Y"
 
-  have "AE x in XY.P. 1 = RN_deriv XY.P (ereal\<circ>joint_distribution X Y) x"
-  proof (rule XY.RN_deriv_unique[OF _ ac])
-    show "measure_space ?J" by default
-    fix A assume "A \<in> sets XY.P"
-    then show "(ereal\<circ>joint_distribution X Y) A = (\<integral>\<^isup>+ x. 1 * indicator A x \<partial>XY.P)"
-      by (simp add: j_eq)
-  qed (insert XY.measurable_const[of 1 borel], auto)
-  then have ae_XY: "AE x in XY.P. entropy_density b XY.P (ereal\<circ>joint_distribution X Y) x = 0"
-    by (elim XY.AE_mp) (simp add: entropy_density_def)
-  have ae_J: "AE x in ?J. entropy_density b XY.P (ereal\<circ>joint_distribution X Y) x = 0"
-  proof (rule XY.absolutely_continuous_AE)
-    show "measure_space ?J" by default
-    show "XY.absolutely_continuous (measure ?J)"
-      using ac by simp
-  qed (insert ae_XY, simp_all)
-  then show "integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>)
-        (entropy_density b P (ereal\<circ>joint_distribution X Y))"
-    unfolding P_def
-    using ed XY.measurable_const[of 0 borel]
-    by (subst J.integrable_cong_AE) auto
+  interpret X: prob_space "distr M S X"
+    by (rule prob_space_distr) fact
+  interpret Y: prob_space "distr M T Y"
+    by (rule prob_space_distr) fact
+  interpret XY: pair_prob_space "distr M S X" "distr M T Y" by default
+  interpret P: information_space P b unfolding P_def by default (rule b_gt_1)
 
-  show "mutual_information b S T X Y = 0"
-    unfolding mutual_information_def KL_divergence_def P_def
-    by (subst J.integral_cong_AE[OF ae_J]) simp
-next
-  assume "sigma_algebra S" "X \<in> measurable M S" "sigma_algebra T" "Y \<in> measurable M T"
-  then have rvs: "random_variable S X" "random_variable T Y" by blast+
+  interpret Q: prob_space Q unfolding Q_def
+    by (rule prob_space_distr) (simp add: comp_def measurable_pair_iff rv)
 
-  interpret X: prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
-    by (rule distribution_prob_space) fact
-  interpret Y: prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
-    by (rule distribution_prob_space) fact
-  interpret XY: pair_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default
-  interpret XY: information_space XY.P b by default (rule b_gt_1)
+  { assume "distr M S X \<Otimes>\<^isub>M distr M T Y = distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+    then have [simp]: "Q = P"  unfolding Q_def P_def by simp
 
-  let ?J = "XY.P\<lparr> measure := (ereal\<circ>joint_distribution X Y) \<rparr>"
-  interpret J: prob_space ?J
-    using rvs by (intro joint_distribution_prob_space) auto
+    show ac: "absolutely_continuous P Q" by (simp add: absolutely_continuous_def)
+    then have ed: "entropy_density b P Q \<in> borel_measurable P"
+      by (rule P.measurable_entropy_density) simp
 
-  assume ac: "measure_space.absolutely_continuous P (ereal\<circ>joint_distribution X Y)"
-  assume int: "integrable (P\<lparr>measure := (ereal\<circ>joint_distribution X Y)\<rparr>)
-        (entropy_density b P (ereal\<circ>joint_distribution X Y))"
-  assume I_eq_0: "mutual_information b S T X Y = 0"
+    have "AE x in P. 1 = RN_deriv P Q x"
+    proof (rule P.RN_deriv_unique)
+      show "density P (\<lambda>x. 1) = Q"
+        unfolding `Q = P` by (intro measure_eqI) (auto simp: emeasure_density)
+    qed auto
+    then have ae_0: "AE x in P. entropy_density b P Q x = 0"
+      by eventually_elim (auto simp: entropy_density_def)
+    then have "integrable P (entropy_density b P Q) \<longleftrightarrow> integrable Q (\<lambda>x. 0)"
+      using ed unfolding `Q = P` by (intro integrable_cong_AE) auto
+    then show "integrable Q (entropy_density b P Q)" by simp
 
-  have eq: "\<forall>A\<in>sets XY.P. (ereal \<circ> joint_distribution X Y) A = XY.\<mu> A"
-  proof (rule XY.KL_eq_0_imp)
-    show "prob_space ?J" by unfold_locales
-    show "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)"
-      using ac by (simp add: P_def)
-    show "integrable ?J (entropy_density b XY.P (ereal\<circ>joint_distribution X Y))"
-      using int by (simp add: P_def)
-    show "KL_divergence b XY.P (ereal\<circ>joint_distribution X Y) = 0"
-      using I_eq_0 unfolding mutual_information_def by (simp add: P_def)
-  qed
-
-  { fix S X assume "sigma_algebra S"
-    interpret S: sigma_algebra S by fact
-    have "Int_stable \<lparr>space = space M, sets = {X -` A \<inter> space M |A. A \<in> sets S}\<rparr>"
-    proof (safe intro!: Int_stableI)
-      fix A B assume "A \<in> sets S" "B \<in> sets S"
-      then show "\<exists>C. (X -` A \<inter> space M) \<inter> (X -` B \<inter> space M) = (X -` C \<inter> space M) \<and> C \<in> sets S"
-        by (intro exI[of _ "A \<inter> B"]) auto
-    qed }
-  note Int_stable = this
+    show "mutual_information b S T X Y = 0"
+      unfolding mutual_information_def KL_divergence_def P_def[symmetric] Q_def[symmetric] `Q = P`
+      using ae_0 by (simp cong: integral_cong_AE) }
 
-  show "indep_var S X T Y" unfolding indep_var_eq
-  proof (intro conjI indep_set_sigma_sets Int_stable)
-    show "indep_set {X -` A \<inter> space M |A. A \<in> sets S} {Y -` A \<inter> space M |A. A \<in> sets T}"
-    proof (safe intro!: indep_setI)
-      { fix A assume "A \<in> sets S" then show "X -` A \<inter> space M \<in> sets M"
-        using `X \<in> measurable M S` by (auto intro: measurable_sets) }
-      { fix A assume "A \<in> sets T" then show "Y -` A \<inter> space M \<in> sets M"
-        using `Y \<in> measurable M T` by (auto intro: measurable_sets) }
-    next
-      fix A B assume ab: "A \<in> sets S" "B \<in> sets T"
-      have "ereal (prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M))) =
-        ereal (joint_distribution X Y (A \<times> B))"
-        unfolding distribution_def
-        by (intro arg_cong[where f="\<lambda>C. ereal (prob C)"]) auto
-      also have "\<dots> = XY.\<mu> (A \<times> B)"
-        using ab eq by (auto simp: XY.finite_measure_eq)
-      also have "\<dots> = ereal (distribution X A) * ereal (distribution Y B)"
-        using ab by (simp add: XY.pair_measure_times)
-      finally show "prob ((X -` A \<inter> space M) \<inter> (Y -` B \<inter> space M)) =
-        prob (X -` A \<inter> space M) * prob (Y -` B \<inter> space M)"
-        unfolding distribution_def by simp
-    qed
-  qed fact+
-qed
+  { assume ac: "absolutely_continuous P Q"
+    assume int: "integrable Q (entropy_density b P Q)"
+    assume I_eq_0: "mutual_information b S T X Y = 0"
 
-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
-    (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>) (ereal\<circ>joint_distribution X Y)"
-  shows "mutual_information b S T X Y = mutual_information b T S Y X"
-proof -
-  let ?S = "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" and ?T = "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
-  interpret S: prob_space ?S using X by (rule distribution_prob_space)
-  interpret T: prob_space ?T using Y by (rule distribution_prob_space)
-  interpret P: pair_prob_space ?S ?T ..
-  interpret Q: pair_prob_space ?T ?S ..
-  show ?thesis
-    unfolding mutual_information_def
-  proof (intro Q.KL_divergence_vimage[OF Q.measure_preserving_swap _ _ _ _ ac b_gt_1])
-    show "(\<lambda>(x,y). (y,x)) \<in> measure_preserving
-      (P.P \<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>) (Q.P \<lparr> measure := ereal\<circ>joint_distribution Y X\<rparr>)"
-      using X Y unfolding measurable_def
-      unfolding measure_preserving_def using P.pair_sigma_algebra_swap_measurable
-      by (auto simp add: space_pair_measure distribution_def intro!: arg_cong[where f=\<mu>'])
-    have "prob_space (P.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)"
-      using X Y by (auto intro!: distribution_prob_space random_variable_pairI)
-    then show "measure_space (P.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)"
-      unfolding prob_space_def finite_measure_def sigma_finite_measure_def by simp
-  qed auto
+    have eq: "Q = P"
+    proof (rule P.KL_eq_0_iff_eq_ac[THEN iffD1])
+      show "prob_space Q" by unfold_locales
+      show "absolutely_continuous P Q" by fact
+      show "integrable Q (entropy_density b P Q)" by fact
+      show "sets Q = sets P" by (simp add: P_def Q_def sets_pair_measure)
+      show "KL_divergence b P Q = 0"
+        using I_eq_0 unfolding mutual_information_def by (simp add: P_def Q_def)
+    qed
+    then show "distr M S X \<Otimes>\<^isub>M distr M T Y = distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+      unfolding P_def Q_def .. }
 qed
 
-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), measure = ereal\<circ>distribution X \<rparr>
-    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr> X Y"
+  "\<I>(X ; Y) \<equiv> mutual_information b (count_space (X`space M)) (count_space (Y`space M)) X Y"
 
-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
-    (S\<lparr>measure := ereal\<circ>distribution X\<rparr> \<Otimes>\<^isub>M T\<lparr>measure := ereal\<circ>distribution Y\<rparr>)
-    (ereal\<circ>joint_distribution X Y)"
+lemma (in information_space)
+  fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
+  assumes "sigma_finite_measure S" "sigma_finite_measure T"
+  assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  defines "f \<equiv> \<lambda>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))"
+  shows mutual_information_distr: "mutual_information b S T X Y = integral\<^isup>L (S \<Otimes>\<^isub>M T) f" (is "?M = ?R")
+    and mutual_information_nonneg: "integrable (S \<Otimes>\<^isub>M T) f \<Longrightarrow> 0 \<le> mutual_information b S T X Y"
 proof -
-  interpret X: finite_prob_space "S\<lparr>measure := ereal\<circ>distribution X\<rparr>"
-    using X by (rule distribution_finite_prob_space)
-  interpret Y: finite_prob_space "T\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
-    using Y by (rule distribution_finite_prob_space)
-  interpret XY: pair_finite_prob_space
-    "S\<lparr>measure := ereal\<circ>distribution X\<rparr>" "T\<lparr> measure := ereal\<circ>distribution Y\<rparr>" by default
-  interpret P: finite_prob_space "XY.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>"
-    using assms by (auto intro!: joint_distribution_finite_prob_space)
-  note rv = assms[THEN finite_random_variableD]
-  show "XY.absolutely_continuous (ereal\<circ>joint_distribution X Y)"
-  proof (rule XY.absolutely_continuousI)
-    show "finite_measure_space (XY.P\<lparr> measure := ereal\<circ>joint_distribution X Y\<rparr>)" by unfold_locales
-    fix x assume "x \<in> space XY.P" and "XY.\<mu> {x} = 0"
-    then obtain a b where "x = (a, b)"
-      and "distribution X {a} = 0 \<or> distribution Y {b} = 0"
-      by (cases x) (auto simp: space_pair_measure)
-    with finite_distribution_order(5,6)[OF X Y]
-    show "(ereal \<circ> joint_distribution X Y) {x} = 0" by auto
+  have X: "random_variable S X"
+    using Px by (auto simp: distributed_def)
+  have Y: "random_variable T Y"
+    using Py by (auto simp: distributed_def)
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T ..
+  interpret X: prob_space "distr M S X" using X by (rule prob_space_distr)
+  interpret Y: prob_space "distr M T Y" using Y by (rule prob_space_distr)
+  interpret XY: pair_prob_space "distr M S X" "distr M T Y" ..
+  let ?P = "S \<Otimes>\<^isub>M T"
+  let ?D = "distr M ?P (\<lambda>x. (X x, Y x))"
+
+  { fix A assume "A \<in> sets S"
+    with X Y have "emeasure (distr M S X) A = emeasure ?D (A \<times> space T)"
+      by (auto simp: emeasure_distr measurable_Pair measurable_space
+               intro!: arg_cong[where f="emeasure M"]) }
+  note marginal_eq1 = this
+  { fix A assume "A \<in> sets T"
+    with X Y have "emeasure (distr M T Y) A = emeasure ?D (space S \<times> A)"
+      by (auto simp: emeasure_distr measurable_Pair measurable_space
+               intro!: arg_cong[where f="emeasure M"]) }
+  note marginal_eq2 = this
+
+  have eq: "(\<lambda>x. ereal (Px (fst x) * Py (snd x))) = (\<lambda>(x, y). ereal (Px x) * ereal (Py y))"
+    by auto
+
+  have distr_eq: "distr M S X \<Otimes>\<^isub>M distr M T Y = density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))"
+    unfolding Px(1)[THEN distributed_distr_eq_density] Py(1)[THEN distributed_distr_eq_density] eq
+  proof (subst pair_measure_density)
+    show "(\<lambda>x. ereal (Px x)) \<in> borel_measurable S" "(\<lambda>y. ereal (Py y)) \<in> borel_measurable T"
+      "AE x in S. 0 \<le> ereal (Px x)" "AE y in T. 0 \<le> ereal (Py y)"
+      using Px Py by (auto simp: distributed_def)
+    show "sigma_finite_measure (density S Px)" unfolding Px(1)[THEN distributed_distr_eq_density, symmetric] ..
+    show "sigma_finite_measure (density T Py)" unfolding Py(1)[THEN distributed_distr_eq_density, symmetric] ..
+  qed (fact | simp)+
+  
+  have M: "?M = KL_divergence b (density ?P (\<lambda>x. ereal (Px (fst x) * Py (snd x)))) (density ?P (\<lambda>x. ereal (Pxy x)))"
+    unfolding mutual_information_def distr_eq Pxy(1)[THEN distributed_distr_eq_density] ..
+
+  from Px Py have f: "(\<lambda>x. Px (fst x) * Py (snd x)) \<in> borel_measurable ?P"
+    by (intro borel_measurable_times) (auto intro: distributed_real_measurable measurable_fst'' measurable_snd'')
+  have PxPy_nonneg: "AE x in ?P. 0 \<le> Px (fst x) * Py (snd x)"
+  proof (rule ST.AE_pair_measure)
+    show "{x \<in> space ?P. 0 \<le> Px (fst x) * Py (snd x)} \<in> sets ?P"
+      using f by auto
+    show "AE x in S. AE y in T. 0 \<le> Px (fst (x, y)) * Py (snd (x, y))"
+      using Px Py by (auto simp: zero_le_mult_iff dest!: distributed_real_AE)
+  qed
+
+  have "(AE x in ?P. Px (fst x) = 0 \<longrightarrow> Pxy x = 0)"
+    by (rule subdensity_real[OF measurable_fst Pxy Px]) auto
+  moreover
+  have "(AE x in ?P. Py (snd x) = 0 \<longrightarrow> Pxy x = 0)"
+    by (rule subdensity_real[OF measurable_snd Pxy Py]) auto
+  ultimately have ac: "AE x in ?P. Px (fst x) * Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+    by eventually_elim auto
+
+  show "?M = ?R"
+    unfolding M f_def
+    using b_gt_1 f PxPy_nonneg Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] ac
+    by (rule ST.KL_density_density)
+
+  assume int: "integrable (S \<Otimes>\<^isub>M T) f"
+  show "0 \<le> ?M" unfolding M
+  proof (rule ST.KL_density_density_nonneg
+    [OF b_gt_1 f PxPy_nonneg _ Pxy[THEN distributed_real_measurable] Pxy[THEN distributed_real_AE] _ ac int[unfolded f_def]])
+    show "prob_space (density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Pxy x))) "
+      unfolding distributed_distr_eq_density[OF Pxy, symmetric]
+      using distributed_measurable[OF Pxy] by (rule prob_space_distr)
+    show "prob_space (density (S \<Otimes>\<^isub>M T) (\<lambda>x. ereal (Px (fst x) * Py (snd x))))"
+      unfolding distr_eq[symmetric] by unfold_locales
   qed
 qed
 
 lemma (in information_space)
-  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.
-      joint_distribution X Y {(x,y)} *
-      log b (joint_distribution X Y {(x,y)} /
-      (distribution X {x} * distribution Y {y})))"
-    (is ?sum)
-  and mutual_information_positive_generic:
-     "0 \<le> mutual_information b MX MY X Y" (is ?positive)
+  fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
+  assumes "sigma_finite_measure S" "sigma_finite_measure T"
+  assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  assumes ae: "AE x in S. AE y in T. Pxy (x, y) = Px x * Py y"
+  shows mutual_information_eq_0: "mutual_information b S T X Y = 0"
 proof -
-  interpret X: finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>"
-    using MX by (rule distribution_finite_prob_space)
-  interpret Y: finite_prob_space "MY\<lparr>measure := ereal\<circ>distribution Y\<rparr>"
-    using MY by (rule distribution_finite_prob_space)
-  interpret XY: pair_finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>" "MY\<lparr>measure := ereal\<circ>distribution Y\<rparr>" by default
-  interpret P: finite_prob_space "XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>"
-    using assms by (auto intro!: joint_distribution_finite_prob_space)
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T ..
 
-  have P_ms: "finite_measure_space (XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>)" by unfold_locales
-  have P_ps: "finite_prob_space (XY.P\<lparr>measure := ereal\<circ>joint_distribution X Y\<rparr>)" by unfold_locales
-
-  show ?sum
-    unfolding Let_def mutual_information_def
-    by (subst XY.KL_divergence_eq_finite[OF P_ms finite_variables_absolutely_continuous[OF MX MY]])
-       (auto simp add: space_pair_measure setsum_cartesian_product')
-
-  show ?positive
-    using XY.KL_divergence_positive_finite[OF P_ps finite_variables_absolutely_continuous[OF MX MY] b_gt_1]
-    unfolding mutual_information_def .
+  have "AE x in S \<Otimes>\<^isub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
+    by (rule subdensity_real[OF measurable_fst Pxy Px]) auto
+  moreover
+  have "AE x in S \<Otimes>\<^isub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+    by (rule subdensity_real[OF measurable_snd Pxy Py]) auto
+  moreover 
+  have "AE x in S \<Otimes>\<^isub>M T. Pxy x = Px (fst x) * Py (snd x)"
+    using distributed_real_measurable[OF Px] distributed_real_measurable[OF Py] distributed_real_measurable[OF Pxy]
+    by (intro ST.AE_pair_measure) (auto simp: ae intro!: measurable_snd'' measurable_fst'')
+  ultimately have "AE x in S \<Otimes>\<^isub>M T. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) = 0"
+    by eventually_elim simp
+  then have "(\<integral>x. Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) \<partial>(S \<Otimes>\<^isub>M T)) = (\<integral>x. 0 \<partial>(S \<Otimes>\<^isub>M T))"
+    by (rule integral_cong_AE)
+  then show ?thesis
+    by (subst mutual_information_distr[OF assms(1-5)]) simp
 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"
-  unfolding mutual_information_generic_eq[OF X Y] mutual_information_generic_eq[OF Y X]
-  unfolding joint_distribution_commute_singleton[of X Y]
-  by (auto simp add: ac_simps intro!: setsum_reindex_cong[OF swap_inj_on])
-
-lemma (in information_space) mutual_information_commute_simple:
-  assumes X: "simple_function M X" and Y: "simple_function M Y"
-  shows "\<I>(X;Y) = \<I>(Y;X)"
-  by (intro mutual_information_commute X Y simple_function_imp_finite_random_variable)
-
-lemma (in information_space) mutual_information_eq:
-  assumes "simple_function M X" "simple_function M Y"
-  shows "\<I>(X;Y) = (\<Sum> (x,y) \<in> X ` space M \<times> Y ` space M.
-    distribution (\<lambda>x. (X x, Y x)) {(x,y)} * log b (distribution (\<lambda>x. (X x, Y x)) {(x,y)} /
-                                                   (distribution X {x} * distribution Y {y})))"
-  using assms by (simp add: mutual_information_generic_eq)
+lemma (in information_space) mutual_information_simple_distributed:
+  assumes X: "simple_distributed M X Px" and Y: "simple_distributed M Y Py"
+  assumes XY: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
+  shows "\<I>(X ; Y) = (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x))`space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))"
+proof (subst mutual_information_distr[OF _ _ simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]])
+  note fin = simple_distributed_joint_finite[OF XY, simp]
+  show "sigma_finite_measure (count_space (X ` space M))"
+    by (simp add: sigma_finite_measure_count_space_finite)
+  show "sigma_finite_measure (count_space (Y ` space M))"
+    by (simp add: sigma_finite_measure_count_space_finite)
+  let ?Pxy = "\<lambda>x. (if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x else 0)"
+  let ?f = "\<lambda>x. ?Pxy x * log b (?Pxy x / (Px (fst x) * Py (snd x)))"
+  have "\<And>x. ?f x = (if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x))) else 0)"
+    by auto
+  with fin show "(\<integral> x. ?f x \<partial>(count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M))) =
+    (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y)))"
+    by (auto simp add: pair_measure_count_space lebesgue_integral_count_space_finite setsum_cases split_beta'
+             intro!: setsum_cong)
+qed
 
-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 "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 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 M X" "simple_function M Y"
-  shows "0 \<le> \<I>(X;Y)"
-  using assms by (simp add: mutual_information_positive_generic)
+lemma (in information_space)
+  fixes Pxy :: "'b \<times> 'c \<Rightarrow> real" and Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
+  assumes Px: "simple_distributed M X Px" and Py: "simple_distributed M Y Py"
+  assumes Pxy: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
+  assumes ae: "\<forall>x\<in>space M. Pxy (X x, Y x) = Px (X x) * Py (Y x)"
+  shows mutual_information_eq_0_simple: "\<I>(X ; Y) = 0"
+proof (subst mutual_information_simple_distributed[OF Px Py Pxy])
+  have "(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) =
+    (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. 0)"
+    by (intro setsum_cong) (auto simp: ae)
+  then show "(\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M.
+    Pxy (x, y) * log b (Pxy (x, y) / (Px x * Py y))) = 0" by simp
+qed
 
 subsection {* Entropy *}
 
+definition (in prob_space) entropy :: "real \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> real" where
+  "entropy b S X = - KL_divergence b S (distr M S X)"
+
 abbreviation (in information_space)
   entropy_Pow ("\<H>'(_')") where
-  "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M), measure = ereal\<circ>distribution X \<rparr> X"
-
-lemma (in information_space) entropy_generic_eq:
-  fixes X :: "'a \<Rightarrow> 'c"
-  assumes MX: "finite_random_variable MX X"
-  shows "entropy b MX X = -(\<Sum> x \<in> space MX. distribution X {x} * log b (distribution X {x}))"
-proof -
-  interpret MX: finite_prob_space "MX\<lparr>measure := ereal\<circ>distribution X\<rparr>"
-    using MX by (rule distribution_finite_prob_space)
-  let ?X = "\<lambda>x. distribution X {x}"
-  let ?XX = "\<lambda>x y. joint_distribution X X {(x, y)}"
+  "\<H>(X) \<equiv> entropy b (count_space (X`space M)) X"
 
-  { fix x y :: 'c
-    { assume "x \<noteq> y"
-      then have "(\<lambda>x. (X x, X x)) -` {(x,y)} \<inter> space M = {}" by auto
-      then have "joint_distribution X X {(x, y)} = 0" by (simp add: distribution_def) }
-    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)"
-      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]
-    unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX
-    using MX.finite_space by (auto simp: setsum_cases)
+lemma (in information_space) entropy_distr:
+  fixes X :: "'a \<Rightarrow> 'b"
+  assumes "sigma_finite_measure MX" and X: "distributed M MX X f"
+  shows "entropy b MX X = - (\<integral>x. f x * log b (f x) \<partial>MX)"
+proof -
+  interpret MX: sigma_finite_measure MX by fact
+  from X show ?thesis
+    unfolding entropy_def X[THEN distributed_distr_eq_density]
+    by (subst MX.KL_density[OF b_gt_1]) (simp_all add: distributed_real_AE distributed_real_measurable)
 qed
 
-lemma (in information_space) entropy_eq:
-  assumes "simple_function M X"
-  shows "\<H>(X) = -(\<Sum> x \<in> X ` space M. distribution X {x} * log b (distribution X {x}))"
-  using assms by (simp add: entropy_generic_eq)
-
-lemma (in information_space) entropy_positive:
-  "simple_function M X \<Longrightarrow> 0 \<le> \<H>(X)"
-  unfolding entropy_def by (simp add: mutual_information_positive)
+lemma (in information_space) entropy_uniform:
+  assumes "sigma_finite_measure MX"
+  assumes A: "A \<in> sets MX" "emeasure MX A \<noteq> 0" "emeasure MX A \<noteq> \<infinity>"
+  assumes X: "distributed M MX X (\<lambda>x. 1 / measure MX A * indicator A x)"
+  shows "entropy b MX X = log b (measure MX A)"
+proof (subst entropy_distr[OF _ X])
+  let ?f = "\<lambda>x. 1 / measure MX A * indicator A x"
+  have "- (\<integral>x. ?f x * log b (?f x) \<partial>MX) = 
+    - (\<integral>x. (log b (1 / measure MX A) / measure MX A) * indicator A x \<partial>MX)"
+    by (auto intro!: integral_cong simp: indicator_def)
+  also have "\<dots> = - log b (inverse (measure MX A))"
+    using A by (subst integral_cmult(2))
+               (simp_all add: measure_def real_of_ereal_eq_0 integral_cmult inverse_eq_divide)
+  also have "\<dots> = log b (measure MX A)"
+    using b_gt_1 A by (subst log_inverse) (auto simp add: measure_def less_le real_of_ereal_eq_0
+                                                          emeasure_nonneg real_of_ereal_pos)
+  finally show "- (\<integral>x. ?f x * log b (?f x) \<partial>MX) = log b (measure MX A)" by simp
+qed fact+
 
-lemma (in information_space) entropy_certainty_eq_0:
-  assumes X: "simple_function M X" and "x \<in> X ` space M" and "distribution X {x} = 1"
-  shows "\<H>(X) = 0"
-proof -
-  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = ereal\<circ>distribution X\<rparr>"
-  note simple_function_imp_finite_random_variable[OF `simple_function M X`]
-  from distribution_finite_prob_space[OF this, of "\<lparr> measure = ereal\<circ>distribution X \<rparr>"]
-  interpret X: finite_prob_space ?X by simp
-  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 *: "y \<in> X ` space M"
-    { assume asm: "y \<noteq> x"
-      with * have "{y} \<subseteq> X ` space M - {x}" by auto
-      from X.measure_mono[OF this] X0 asm *
-      have "distribution X {y} = 0"  by (auto intro: antisym) }
-    then have "distribution X {y} = (if x = y then 1 else 0)"
-      using assms by auto }
-  note fi = this
-  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[OF `simple_function M X`] by (auto simp: y fi)
+lemma (in information_space) entropy_simple_distributed:
+  fixes X :: "'a \<Rightarrow> 'b"
+  assumes X: "simple_distributed M X f"
+  shows "\<H>(X) = - (\<Sum>x\<in>X`space M. f x * log b (f x))"
+proof (subst entropy_distr[OF _ simple_distributed[OF X]])
+  show "sigma_finite_measure (count_space (X ` space M))"
+    using X by (simp add: sigma_finite_measure_count_space_finite simple_distributed_def)
+  show "- (\<integral>x. f x * log b (f x) \<partial>(count_space (X`space M))) = - (\<Sum>x\<in>X ` space M. f x * log b (f x))"
+    using X by (auto simp add: lebesgue_integral_count_space_finite)
 qed
 
 lemma (in information_space) entropy_le_card_not_0:
-  assumes X: "simple_function M X"
-  shows "\<H>(X) \<le> log b (card (X ` space M \<inter> {x. distribution X {x} \<noteq> 0}))"
+  assumes X: "simple_distributed M X f"
+  shows "\<H>(X) \<le> log b (card (X ` space M \<inter> {x. f x \<noteq> 0}))"
 proof -
-  let ?p = "\<lambda>x. distribution X {x}"
-  have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))"
-    unfolding entropy_eq[OF X] setsum_negf[symmetric]
-    by (auto intro!: setsum_cong simp: log_simps)
-  also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))"
-    using not_empty b_gt_1 `simple_function M X` sum_over_space_real_distribution[OF X]
-    by (intro log_setsum') (auto simp: simple_function_def)
-  also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?p x \<noteq> 0 then 1 else 0)"
+  have "\<H>(X) = (\<Sum>x\<in>X`space M. f x * log b (1 / f x))"
+    unfolding entropy_simple_distributed[OF X] setsum_negf[symmetric]
+    using X by (auto dest: simple_distributed_nonneg intro!: setsum_cong simp: log_simps less_le)
+  also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. f x * (1 / f x))"
+    using not_empty b_gt_1 `simple_distributed M X f`
+    by (intro log_setsum') (auto simp: simple_distributed_nonneg simple_distributed_setsum_space)
+  also have "\<dots> = log b (\<Sum>x\<in>X`space M. if f x \<noteq> 0 then 1 else 0)"
     by (intro arg_cong[where f="\<lambda>X. log b X"] setsum_cong) auto
   finally show ?thesis
-    using `simple_function M X` by (auto simp: setsum_cases real_eq_of_nat simple_function_def)
-qed
-
-lemma (in prob_space) measure'_translate:
-  assumes X: "random_variable S X" and A: "A \<in> sets S"
-  shows "finite_measure.\<mu>' (S\<lparr> measure := ereal\<circ>distribution X \<rparr>) A = distribution X A"
-proof -
-  interpret S: prob_space "S\<lparr> measure := ereal\<circ>distribution X \<rparr>"
-    using distribution_prob_space[OF X] .
-  from A show "S.\<mu>' A = distribution X A"
-    unfolding S.\<mu>'_def by (simp add: distribution_def [abs_def] \<mu>'_def)
-qed
-
-lemma (in information_space) entropy_uniform_max:
-  assumes X: "simple_function M 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 -
-  let ?X = "\<lparr> space = X ` space M, sets = Pow (X ` space M), measure = undefined\<rparr>\<lparr> measure := ereal\<circ>distribution X\<rparr>"
-  note frv = simple_function_imp_finite_random_variable[OF X]
-  from distribution_finite_prob_space[OF this, of "\<lparr> measure = ereal\<circ>distribution X \<rparr>"]
-  interpret X: finite_prob_space ?X by simp
-  note rv = finite_random_variableD[OF frv]
-  have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
-    using `simple_function M X` not_empty by (auto simp: simple_function_def)
-  { fix x assume "x \<in> space ?X"
-    moreover then have "X.\<mu>' {x} = 1 / card (space ?X)"
-    proof (rule X.uniform_prob)
-      fix x y assume "x \<in> space ?X" "y \<in> space ?X"
-      with assms(2)[of x y] show "X.\<mu>' {x} = X.\<mu>' {y}"
-        by (subst (1 2) measure'_translate[OF rv]) auto
-    qed
-    ultimately have "distribution X {x} = 1 / card (space ?X)"
-      by (subst (asm) measure'_translate[OF rv]) auto }
-  thus ?thesis
-    using not_empty X.finite_space b_gt_1 card_gt0
-    by (simp add: entropy_eq[OF `simple_function M X`] real_eq_of_nat[symmetric] log_simps)
+    using `simple_distributed M X f` by (auto simp: setsum_cases real_eq_of_nat)
 qed
 
 lemma (in information_space) entropy_le_card:
-  assumes "simple_function M X"
+  assumes "simple_distributed M X f"
   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
+  assume "X ` space M \<inter> {x. f x \<noteq> 0} = {}"
+  then have "\<And>x. x\<in>X`space M \<Longrightarrow> f x = 0" by auto
   moreover
   have "0 < card (X`space M)"
-    using `simple_function M X` not_empty
-    by (auto simp: card_gt_0_iff simple_function_def)
+    using `simple_distributed M X f` not_empty by (auto simp: card_gt_0_iff)
   then have "log b 1 \<le> log b (real (card (X`space M)))"
     using b_gt_1 by (intro log_le) auto
-  ultimately show ?thesis using assms by (simp add: entropy_eq)
+  ultimately show ?thesis using assms by (simp add: entropy_simple_distributed)
 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 assms not_empty by (auto intro!: card_mono simp: simple_function_def)
+  assume False: "X ` space M \<inter> {x. f x \<noteq> 0} \<noteq> {}"
+  have "card (X ` space M \<inter> {x. f x \<noteq> 0}) \<le> card (X ` space M)"
+    (is "?A \<le> ?B") using assms not_empty
+    by (auto intro!: card_mono simp: simple_function_def simple_distributed_def)
   note entropy_le_card_not_0[OF assms]
   also have "log b (real ?A) \<le> log b (real ?B)"
     using b_gt_1 False not_empty `?A \<le> ?B` assms
-    by (auto intro!: log_le simp: card_gt_0_iff simp: simple_function_def)
+    by (auto intro!: log_le simp: card_gt_0_iff simp: simple_distributed_def)
   finally show ?thesis .
 qed
 
-lemma (in information_space) entropy_commute:
-  assumes "simple_function M X" "simple_function M Y"
-  shows "\<H>(\<lambda>x. (X x, Y x)) = \<H>(\<lambda>x. (Y x, X x))"
-proof -
-  have sf: "simple_function M (\<lambda>x. (X x, Y x))" "simple_function M (\<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 sf[THEN entropy_eq] unfolding * setsum_reindex[OF inj]
-    by (simp add: joint_distribution_commute[of Y X] split_beta)
-qed
-
-lemma (in information_space) entropy_eq_cartesian_product:
-  assumes "simple_function M X" "simple_function M Y"
-  shows "\<H>(\<lambda>x. (X x, Y x)) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
-    joint_distribution X Y {(x,y)} * log b (joint_distribution X Y {(x,y)}))"
-proof -
-  have sf: "simple_function M (\<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 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)
@@ -917,489 +755,553 @@
 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), measure = ereal\<circ>distribution X \<rparr>
-    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr>
-    \<lparr> space = Z`space M, sets = Pow (Z`space M), measure = ereal\<circ>distribution Z \<rparr>
-    X Y Z"
+    (count_space (X ` space M)) (count_space (Y ` space M)) (count_space (Z ` space M)) X Y Z"
 
 lemma (in information_space) conditional_mutual_information_generic_eq:
-  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.
-             distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
-             log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} /
-    (joint_distribution X Z {(x, z)} * (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
-  (is "_ = (\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z z))))")
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T" and P: "sigma_finite_measure P"
+  assumes Px: "distributed M S X Px"
+  assumes Pz: "distributed M P Z Pz"
+  assumes Pyz: "distributed M (T \<Otimes>\<^isub>M P) (\<lambda>x. (Y x, Z x)) Pyz"
+  assumes Pxz: "distributed M (S \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Z x)) Pxz"
+  assumes Pxyz: "distributed M (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Y x, Z x)) Pxyz"
+  assumes I1: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Px x * Pyz (y, z))))"
+  assumes I2: "integrable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). Pxyz (x, y, z) * log b (Pxz (x, z) / (Px x * Pz z)))"
+  shows "conditional_mutual_information b S T P X Y Z
+    = (\<integral>(x, y, z). Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
 proof -
-  let ?X = "\<lambda>x. distribution X {x}"
-  note finite_var = MX MY MZ
-  note YZ = finite_random_variable_pairI[OF finite_var(2,3)]
-  note XYZ = finite_random_variable_pairI[OF MX YZ]
-  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]
-    finite_distribution_order(5,6)[OF finite_var(1,3)]
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret P: sigma_finite_measure P by fact
+  interpret TP: pair_sigma_finite T P ..
+  interpret SP: pair_sigma_finite S P ..
+  interpret SPT: pair_sigma_finite "S \<Otimes>\<^isub>M P" T ..
+  interpret STP: pair_sigma_finite S "T \<Otimes>\<^isub>M P" ..
+  have TP: "sigma_finite_measure (T \<Otimes>\<^isub>M P)" ..
+  have SP: "sigma_finite_measure (S \<Otimes>\<^isub>M P)" ..
+  have YZ: "random_variable (T \<Otimes>\<^isub>M P) (\<lambda>x. (Y x, Z x))"
+    using Pyz by (simp add: distributed_measurable)
+
+  have Pxyz_f: "\<And>M f. f \<in> measurable M (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) \<Longrightarrow> (\<lambda>x. Pxyz (f x)) \<in> borel_measurable M"
+    using measurable_comp[OF _ Pxyz[THEN distributed_real_measurable]] by (auto simp: comp_def)
+
+  { fix f g h M
+    assume f: "f \<in> measurable M S" and g: "g \<in> measurable M P" and h: "h \<in> measurable M (S \<Otimes>\<^isub>M P)"
+    from measurable_comp[OF h Pxz[THEN distributed_real_measurable]]
+         measurable_comp[OF f Px[THEN distributed_real_measurable]]
+         measurable_comp[OF g Pz[THEN distributed_real_measurable]]
+    have "(\<lambda>x. log b (Pxz (h x) / (Px (f x) * Pz (g x)))) \<in> borel_measurable M"
+      by (simp add: comp_def b_gt_1) }
+  note borel_log = this
+
+  have measurable_cut: "(\<lambda>(x, y, z). (x, z)) \<in> measurable (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (S \<Otimes>\<^isub>M P)"
+    by (auto simp add: split_beta' comp_def intro!: measurable_Pair measurable_snd')
+  
+  from Pxz Pxyz have distr_eq: "distr M (S \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Z x)) =
+    distr (distr M (S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P) (\<lambda>x. (X x, Y x, Z x))) (S \<Otimes>\<^isub>M P) (\<lambda>(x, y, z). (x, z))"
+    by (subst distr_distr[OF measurable_cut]) (auto dest: distributed_measurable simp: comp_def)
 
-  note random_var = finite_var[THEN finite_random_variableD]
-  note finite = finite_var(1) YZ finite_var(3) XZ YZX
-
-  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
-
-  have "(\<Sum>(x, y, z)\<in>?S. ?XYZ x y z * log b (?XYZ x y z / (?XZ x z * (?YZ y z / ?Z 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"
-    show "?L x y z = ?R x y z"
+  have "mutual_information b S P X Z =
+    (\<integral>x. Pxz x * log b (Pxz x / (Px (fst x) * Pz (snd x))) \<partial>(S \<Otimes>\<^isub>M P))"
+    by (rule mutual_information_distr[OF S P Px Pz Pxz])
+  also have "\<dots> = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))"
+    using b_gt_1 Pxz Px Pz
+    by (subst distributed_transform_integral[OF Pxyz Pxz, where T="\<lambda>(x, y, z). (x, z)"])
+       (auto simp: split_beta' intro!: measurable_Pair measurable_snd' measurable_snd'' measurable_fst'' borel_measurable_times
+             dest!: distributed_real_measurable)
+  finally have mi_eq:
+    "mutual_information b S P X Z = (\<integral>(x,y,z). Pxyz (x,y,z) * log b (Pxz (x,z) / (Px x * Pz z)) \<partial>(S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P))" .
+  
+  have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Px (fst x) = 0 \<longrightarrow> Pxyz x = 0"
+    by (intro subdensity_real[of fst, OF _ Pxyz Px]) auto
+  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+    by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz Pz]) (auto intro: measurable_snd')
+  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+    by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz Pxz]) (auto intro: measurable_Pair measurable_snd')
+  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
+    by (intro subdensity_real[of snd, OF _ Pxyz Pyz]) (auto intro: measurable_Pair)
+  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Px (fst x)"
+    using Px by (intro STP.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable)
+  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pyz (snd x)"
+    using Pyz by (intro STP.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
+  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pz (snd (snd x))"
+    using Pz Pz[THEN distributed_real_measurable] by (auto intro!: measurable_snd'' TP.AE_pair_measure STP.AE_pair_measure AE_I2[of S] dest: distributed_real_AE)
+  moreover have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P. 0 \<le> Pxz (fst x, snd (snd x))"
+    using Pxz[THEN distributed_real_AE, THEN SP.AE_pair]
+    using measurable_comp[OF measurable_Pair[OF measurable_fst measurable_comp[OF measurable_snd measurable_snd]] Pxz[THEN distributed_real_measurable], of T]
+    using measurable_comp[OF measurable_snd measurable_Pair2[OF Pxz[THEN distributed_real_measurable]], of _ T]
+    by (auto intro!: TP.AE_pair_measure STP.AE_pair_measure simp: comp_def)
+  moreover note Pxyz[THEN distributed_real_AE]
+  ultimately have "AE x in S \<Otimes>\<^isub>M T \<Otimes>\<^isub>M P.
+    Pxyz x * log b (Pxyz x / (Px (fst x) * Pyz (snd x))) -
+    Pxyz x * log b (Pxz (fst x, snd (snd x)) / (Px (fst x) * Pz (snd (snd x)))) =
+    Pxyz x * log b (Pxyz x * Pz (snd (snd x)) / (Pxz (fst x, snd (snd x)) * Pyz (snd x))) "
+  proof eventually_elim
+    case (goal1 x)
+    show ?case
     proof cases
-      assume "?XYZ x y z \<noteq> 0"
-      with space have "0 < ?X x" "0 < ?Z z" "0 < ?XZ x z" "0 < ?YZ y z" "0 < ?XYZ x y z"
-        using order1 order2 by (auto simp: less_le)
-      with b_gt_1 show ?thesis
-        by (simp add: log_mult log_divide zero_less_mult_iff zero_less_divide_iff)
+      assume "Pxyz x \<noteq> 0"
+      with goal1 have "0 < Px (fst x)" "0 < Pz (snd (snd x))" "0 < Pxz (fst x, snd (snd x))" "0 < Pyz (snd x)" "0 < Pxyz x"
+        by auto
+      then show ?thesis
+        using b_gt_1 by (simp add: log_simps mult_pos_pos less_imp_le field_simps)
     qed simp
   qed
-  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_joint_distribution_singleton[OF finite_var(2) ZX]
-    by (intro setsum_cong refl) (simp add: space_pair_measure)
-  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"
+  with I1 I2 show ?thesis
     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_measure setsum_cartesian_product')
-  finally show ?thesis by simp
+    apply (subst mi_eq)
+    apply (subst mutual_information_distr[OF S TP Px Pyz Pxyz])
+    apply (subst integral_diff(2)[symmetric])
+    apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff)
+    done
 qed
 
 lemma (in information_space) conditional_mutual_information_eq:
-  assumes "simple_function M X" "simple_function M Y" "simple_function M Z"
-  shows "\<I>(X;Y|Z) = (\<Sum>(x, y, z) \<in> X`space M \<times> Y`space M \<times> Z`space M.
-             distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} *
-             log b (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)} /
-    (joint_distribution X Z {(x, z)} * joint_distribution Y Z {(y,z)} / distribution Z {z})))"
-  by (subst conditional_mutual_information_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]])
-     simp
+  assumes Pz: "simple_distributed M Z Pz"
+  assumes Pyz: "simple_distributed M (\<lambda>x. (Y x, Z x)) Pyz"
+  assumes Pxz: "simple_distributed M (\<lambda>x. (X x, Z x)) Pxz"
+  assumes Pxyz: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Pxyz"
+  shows "\<I>(X ; Y | Z) =
+   (\<Sum>(x, y, z)\<in>(\<lambda>x. (X x, Y x, Z x))`space M. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
+proof (subst conditional_mutual_information_generic_eq[OF _ _ _ _
+    simple_distributed[OF Pz] simple_distributed_joint[OF Pyz] simple_distributed_joint[OF Pxz]
+    simple_distributed_joint2[OF Pxyz]])
+  note simple_distributed_joint2_finite[OF Pxyz, simp]
+  show "sigma_finite_measure (count_space (X ` space M))"
+    by (simp add: sigma_finite_measure_count_space_finite)
+  show "sigma_finite_measure (count_space (Y ` space M))"
+    by (simp add: sigma_finite_measure_count_space_finite)
+  show "sigma_finite_measure (count_space (Z ` space M))"
+    by (simp add: sigma_finite_measure_count_space_finite)
+  have "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) \<Otimes>\<^isub>M count_space (Z ` space M) =
+      count_space (X`space M \<times> Y`space M \<times> Z`space M)"
+    (is "?P = ?C")
+    by (simp add: pair_measure_count_space)
 
-lemma (in information_space) conditional_mutual_information_eq_mutual_information:
-  assumes X: "simple_function M X" and Y: "simple_function M 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 M (\<lambda>x. ())" by auto
-  show ?thesis
-    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)
+  let ?Px = "\<lambda>x. measure M (X -` {x} \<inter> space M)"
+  have "(\<lambda>x. (X x, Z x)) \<in> measurable M (count_space (X ` space M) \<Otimes>\<^isub>M count_space (Z ` space M))"
+    using simple_distributed_joint[OF Pxz] by (rule distributed_measurable)
+  from measurable_comp[OF this measurable_fst]
+  have "random_variable (count_space (X ` space M)) X"
+    by (simp add: comp_def)
+  then have "simple_function M X"    
+    unfolding simple_function_def by auto
+  then have "simple_distributed M X ?Px"
+    by (rule simple_distributedI) auto
+  then show "distributed M (count_space (X ` space M)) X ?Px"
+    by (rule simple_distributed)
+
+  let ?f = "(\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x)) ` space M then Pxyz x else 0)"
+  let ?g = "(\<lambda>x. if x \<in> (\<lambda>x. (Y x, Z x)) ` space M then Pyz x else 0)"
+  let ?h = "(\<lambda>x. if x \<in> (\<lambda>x. (X x, Z x)) ` space M then Pxz x else 0)"
+  show
+      "integrable ?P (\<lambda>(x, y, z). ?f (x, y, z) * log b (?f (x, y, z) / (?Px x * ?g (y, z))))"
+      "integrable ?P (\<lambda>(x, y, z). ?f (x, y, z) * log b (?h (x, z) / (?Px x * Pz z)))"
+    by (auto intro!: integrable_count_space simp: pair_measure_count_space)
+  let ?i = "\<lambda>x y z. ?f (x, y, z) * log b (?f (x, y, z) / (?h (x, z) * (?g (y, z) / Pz z)))"
+  let ?j = "\<lambda>x y z. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z)))"
+  have "(\<lambda>(x, y, z). ?i x y z) = (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x)) ` space M then ?j (fst x) (fst (snd x)) (snd (snd x)) else 0)"
+    by (auto intro!: ext)
+  then show "(\<integral> (x, y, z). ?i x y z \<partial>?P) = (\<Sum>(x, y, z)\<in>(\<lambda>x. (X x, Y x, Z x)) ` space M. ?j x y z)"
+    by (auto intro!: setsum_cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite setsum_cases split_beta')
 qed
 
-lemma (in information_space) conditional_mutual_information_generic_positive:
-  assumes X: "finite_random_variable MX X" and Y: "finite_random_variable MY Y" and Z: "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 = "distribution (\<lambda>x. (X x, Y x, Z x))"
-  let ?dXZ = "joint_distribution X Z"
-  let ?dYZ = "joint_distribution Y Z"
-  let ?dX = "distribution X"
-  let ?dZ = "distribution Z"
-  let ?M = "space MX \<times> space MY \<times> space MZ"
+lemma (in information_space) conditional_mutual_information_nonneg:
+  assumes X: "simple_function M X" and Y: "simple_function M Y" and Z: "simple_function M Z"
+  shows "0 \<le> \<I>(X ; Y | Z)"
+proof -
+  def Pz \<equiv> "\<lambda>x. if x \<in> Z`space M then measure M (Z -` {x} \<inter> space M) else 0"
+  def Pxz \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (X x, Z x))`space M then measure M ((\<lambda>x. (X x, Z x)) -` {x} \<inter> space M) else 0"
+  def Pyz \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (Y x, Z x))`space M then measure M ((\<lambda>x. (Y x, Z x)) -` {x} \<inter> space M) else 0"
+  def Pxyz \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x))`space M then measure M ((\<lambda>x. (X x, Y x, Z x)) -` {x} \<inter> space M) else 0"
+  let ?M = "X`space M \<times> Y`space M \<times> Z`space M"
 
-  note YZ = finite_random_variable_pairI[OF Y Z]
-  note XZ = finite_random_variable_pairI[OF X Z]
-  note ZX = finite_random_variable_pairI[OF Z X]
-  note YZ = finite_random_variable_pairI[OF Y Z]
-  note XYZ = finite_random_variable_pairI[OF X YZ]
-  note finite = Z YZ XZ XYZ
-  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 Y ZX]
-    by auto
+  note XZ = simple_function_Pair[OF X Z]
+  note YZ = simple_function_Pair[OF Y Z]
+  note XYZ = simple_function_Pair[OF X simple_function_Pair[OF Y Z]]
+  have Pz: "simple_distributed M Z Pz"
+    using Z by (rule simple_distributedI) (auto simp: Pz_def)
+  have Pxz: "simple_distributed M (\<lambda>x. (X x, Z x)) Pxz"
+    using XZ by (rule simple_distributedI) (auto simp: Pxz_def)
+  have Pyz: "simple_distributed M (\<lambda>x. (Y x, Z x)) Pyz"
+    using YZ by (rule simple_distributedI) (auto simp: Pyz_def)
+  have Pxyz: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Pxyz"
+    using XYZ by (rule simple_distributedI) (auto simp: Pxyz_def)
 
-  note order = order
-    finite_distribution_order(5,6)[OF X YZ]
-    finite_distribution_order(5,6)[OF Y Z]
+  { fix z assume z: "z \<in> Z ` space M" then have "(\<Sum>x\<in>X ` space M. Pxz (x, z)) = Pz z"
+      using distributed_marginal_eq_joint_simple[OF X Pz Pxz z]
+      by (auto intro!: setsum_cong simp: Pxz_def) }
+  note marginal1 = this
 
-  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 auto
-  also have "\<dots> \<le> log b (\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z})"
+  { fix z assume z: "z \<in> Z ` space M" then have "(\<Sum>y\<in>Y ` space M. Pyz (y, z)) = Pz z"
+      using distributed_marginal_eq_joint_simple[OF Y Pz Pyz z]
+      by (auto intro!: setsum_cong simp: Pyz_def) }
+  note marginal2 = this
+
+  have "- \<I>(X ; Y | Z) = - (\<Sum>(x, y, z) \<in> ?M. Pxyz (x, y, z) * log b (Pxyz (x, y, z) / (Pxz (x, z) * (Pyz (y,z) / Pz z))))"
+    unfolding conditional_mutual_information_eq[OF Pz Pyz Pxz Pxyz]
+    using X Y Z by (auto intro!: setsum_mono_zero_left simp: Pxyz_def simple_functionD)
+  also have "\<dots> \<le> log b (\<Sum>(x, y, z) \<in> ?M. Pxz (x, z) * (Pyz (y,z) / Pz z))"
     unfolding split_beta'
   proof (rule log_setsum_divide)
-    show "?M \<noteq> {}" using False by simp
+    show "?M \<noteq> {}" using not_empty 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_measure
-                        setsum_joint_distribution_singleton[OF X YZ]
-                        setsum_joint_distribution_singleton[OF Y Z]
-                        setsum_distribution[OF Z])
+    show "finite ?M" using X Y Z by (auto simp: simple_functionD)
 
-    fix x assume "x \<in> ?M"
-    let ?x = "(fst x, fst (snd x), snd (snd x))"
-
-    show "0 \<le> ?dXYZ {?x}"
-      "0 \<le> ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
-     by (simp_all add: mult_nonneg_nonneg divide_nonneg_nonneg)
-
-    assume *: "0 < ?dXYZ {?x}"
-    with `x \<in> ?M` finite order show "0 < ?dXZ {(fst x, snd (snd x))} * ?dYZ {(fst (snd x), snd (snd x))} / ?dZ {snd (snd x)}"
-      by (cases x) (auto simp add: zero_le_mult_iff zero_le_divide_iff less_le)
+    then show "(\<Sum>x\<in>?M. Pxyz (fst x, fst (snd x), snd (snd x))) = 1"
+      apply (subst Pxyz[THEN simple_distributed_setsum_space, symmetric])
+      apply simp
+      apply (intro setsum_mono_zero_right)
+      apply (auto simp: Pxyz_def)
+      done
+    let ?N = "(\<lambda>x. (X x, Y x, Z x)) ` space M"
+    fix x assume x: "x \<in> ?M"
+    let ?Q = "Pxyz (fst x, fst (snd x), snd (snd x))"
+    let ?P = "Pxz (fst x, snd (snd x)) * (Pyz (fst (snd x), snd (snd x)) / Pz (snd (snd x)))"
+    from x show "0 \<le> ?Q" "0 \<le> ?P"
+      using Pxyz[THEN simple_distributed, THEN distributed_real_AE]
+      using Pxz[THEN simple_distributed, THEN distributed_real_AE]
+      using Pyz[THEN simple_distributed, THEN distributed_real_AE]
+      using Pz[THEN simple_distributed, THEN distributed_real_AE]
+      by (auto intro!: mult_nonneg_nonneg divide_nonneg_nonneg simp: AE_count_space Pxyz_def Pxz_def Pyz_def Pz_def)
+    moreover assume "0 < ?Q"
+    moreover have "AE x in count_space ?N. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+      by (intro subdensity_real[of "\<lambda>x. snd (snd x)", OF _ Pxyz[THEN simple_distributed] Pz[THEN simple_distributed]]) (auto intro: measurable_snd')
+    then have "\<And>x. Pz (snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+      by (auto simp: Pz_def Pxyz_def AE_count_space)
+    moreover have "AE x in count_space ?N. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+      by (intro subdensity_real[of "\<lambda>x. (fst x, snd (snd x))", OF _ Pxyz[THEN simple_distributed] Pxz[THEN simple_distributed]]) (auto intro: measurable_Pair measurable_snd')
+    then have "\<And>x. Pxz (fst x, snd (snd x)) = 0 \<longrightarrow> Pxyz x = 0"
+      by (auto simp: Pz_def Pxyz_def AE_count_space)
+    moreover have "AE x in count_space ?N. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
+      by (intro subdensity_real[of snd, OF _ Pxyz[THEN simple_distributed] Pyz[THEN simple_distributed]]) (auto intro: measurable_Pair)
+    then have "\<And>x. Pyz (snd x) = 0 \<longrightarrow> Pxyz x = 0"
+      by (auto simp: Pz_def Pxyz_def AE_count_space)
+    ultimately show "0 < ?P" using x by (auto intro!: divide_pos_pos mult_pos_pos simp: less_le)
   qed
-  also have "(\<Sum>(x, y, z) \<in> ?M. ?dXZ {(x, z)} * ?dYZ {(y,z)} / ?dZ {z}) = (\<Sum>z\<in>space MZ. ?dZ {z})"
+  also have "(\<Sum>(x, y, z) \<in> ?M. Pxz (x, z) * (Pyz (y,z) / Pz z)) = (\<Sum>z\<in>Z`space M. Pz z)"
     apply (simp add: setsum_cartesian_product')
     apply (subst setsum_commute)
     apply (subst (2) setsum_commute)
-    by (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric]
-                   setsum_joint_distribution_singleton[OF X Z]
-                   setsum_joint_distribution_singleton[OF Y Z]
+    apply (auto simp: setsum_divide_distrib[symmetric] setsum_product[symmetric] marginal1 marginal2
           intro!: setsum_cong)
-  also have "log b (\<Sum>z\<in>space MZ. ?dZ {z}) = 0"
-    unfolding setsum_distribution[OF Z] by simp
+    done
+  also have "log b (\<Sum>z\<in>Z`space M. Pz z) = 0"
+    using Pz[THEN simple_distributed_setsum_space] by simp
   finally show ?thesis by simp
 qed
 
-lemma (in information_space) conditional_mutual_information_positive:
-  assumes "simple_function M X" and "simple_function M Y" and "simple_function M Z"
-  shows "0 \<le> \<I>(X;Y|Z)"
-  by (rule conditional_mutual_information_generic_positive[OF assms[THEN simple_function_imp_finite_random_variable]])
-
 subsection {* Conditional Entropy *}
 
 definition (in prob_space)
-  "conditional_entropy b S T X Y = conditional_mutual_information b S S T X X Y"
+  "conditional_entropy b S T X Y = entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) - entropy b T Y"
 
 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), measure = ereal\<circ>distribution X \<rparr>
-    \<lparr> space = Y`space M, sets = Pow (Y`space M), measure = ereal\<circ>distribution Y \<rparr> X Y"
-
-lemma (in information_space) conditional_entropy_positive:
-  "simple_function M X \<Longrightarrow> simple_function M Y \<Longrightarrow> 0 \<le> \<H>(X | Y)"
-  unfolding conditional_entropy_def by (auto intro!: conditional_mutual_information_positive)
+  "\<H>(X | Y) \<equiv> conditional_entropy b (count_space (X`space M)) (count_space (Y`space M)) X Y"
 
 lemma (in information_space) conditional_entropy_generic_eq:
-  fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
-  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.
-         joint_distribution X Z {(x, z)} * log b (joint_distribution X Z {(x, z)} / distribution Z {z}))"
+  fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real"
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
+  assumes Px: "distributed M S X Px"
+  assumes Py: "distributed M T Y Py"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  assumes I1: "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
+  assumes I2: "integrable (S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
+  shows "conditional_entropy b S T X Y = - (\<integral>(x, y). Pxy (x, y) * log b (Pxy (x, y) / Py y) \<partial>(S \<Otimes>\<^isub>M T))"
 proof -
-  interpret MX: finite_sigma_algebra MX using MX by simp
-  interpret MZ: finite_sigma_algebra MZ using MZ by simp
-  let ?XXZ = "\<lambda>x y z. joint_distribution X (\<lambda>x. (X x, Z x)) {(x, y, z)}"
-  let ?XZ = "\<lambda>x z. joint_distribution X Z {(x, z)}"
-  let ?Z = "\<lambda>z. distribution Z {z}"
-  let ?f = "\<lambda>x y z. log b (?XXZ x y z * ?Z z / (?XZ x z * ?XZ y 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' :: 'c and z assume "x' \<noteq> x"
-    then have "?XXZ x x' z = 0"
-      by (auto simp: distribution_def empty_measure'[symmetric]
-               simp del: empty_measure' intro!: arg_cong[where f=\<mu>']) }
-  note this[simp]
-  { fix x x' z assume *: "x \<in> space MX" "z \<in> space MZ"
-    then have "(\<Sum>x'\<in>space MX. ?XXZ x x' z * ?f x x' z)
-      = (\<Sum>x'\<in>space MX. if x = x' then ?XZ x z * ?f x x z else 0)"
-      by (auto intro!: setsum_cong)
-    also have "\<dots> = ?XZ x z * ?f x x z"
-      using `x \<in> space MX` by (simp add: setsum_cases[OF MX.finite_space])
-    also have "\<dots> = ?XZ x z * log b (?Z z / ?XZ x z)" by auto
-    also have "\<dots> = - ?XZ x z * log b (?XZ x z / ?Z z)"
-      using finite_distribution_order(6)[OF MX MZ]
-      by (auto simp: log_simps field_simps zero_less_mult_iff)
-    finally have "(\<Sum>x'\<in>space MX. ?XXZ x x' z * ?f x x' z) = - ?XZ x z * log b (?XZ x z / ?Z z)" . }
-  note * = this
-  show ?thesis
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T ..
+  have ST: "sigma_finite_measure (S \<Otimes>\<^isub>M T)" ..
+
+  interpret Pxy: prob_space "density (S \<Otimes>\<^isub>M T) Pxy"
+    unfolding Pxy[THEN distributed_distr_eq_density, symmetric]
+    using Pxy[THEN distributed_measurable] by (rule prob_space_distr)
+
+  from Py Pxy have distr_eq: "distr M T Y =
+    distr (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) T snd"
+    by (subst distr_distr[OF measurable_snd]) (auto dest: distributed_measurable simp: comp_def)
+
+  have "entropy b T Y = - (\<integral>y. Py y * log b (Py y) \<partial>T)"
+    by (rule entropy_distr[OF T Py])
+  also have "\<dots> = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^isub>M T))"
+    using b_gt_1 Py[THEN distributed_real_measurable]
+    by (subst distributed_transform_integral[OF Pxy Py, where T=snd]) (auto intro!: integral_cong)
+  finally have e_eq: "entropy b T Y = - (\<integral>(x,y). Pxy (x,y) * log b (Py y) \<partial>(S \<Otimes>\<^isub>M T))" .
+  
+  have "AE x in S \<Otimes>\<^isub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
+    by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair)
+  moreover have "AE x in S \<Otimes>\<^isub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+    by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair)
+  moreover have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Px (fst x)"
+    using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable)
+  moreover have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Py (snd x)"
+    using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
+  moreover note Pxy[THEN distributed_real_AE]
+  ultimately have pos: "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Pxy x \<and> 0 \<le> Px (fst x) \<and> 0 \<le> Py (snd x) \<and>
+    (Pxy x = 0 \<or> (Pxy x \<noteq> 0 \<longrightarrow> 0 < Pxy x \<and> 0 < Px (fst x) \<and> 0 < Py (snd x)))"
+    by eventually_elim auto
+
+  from pos have "AE x in S \<Otimes>\<^isub>M T.
+     Pxy x * log b (Pxy x) - Pxy x * log b (Py (snd x)) = Pxy x * log b (Pxy x / Py (snd x))"
+    by eventually_elim (auto simp: log_simps mult_pos_pos field_simps b_gt_1)
+  with I1 I2 show ?thesis
     unfolding conditional_entropy_def
-    unfolding conditional_mutual_information_generic_eq[OF MX MX MZ]
-    by (auto simp: setsum_cartesian_product' setsum_negf[symmetric]
-                   setsum_commute[of _ "space MZ"] *
-             intro!: setsum_cong)
+    apply (subst e_eq)
+    apply (subst entropy_distr[OF ST Pxy])
+    unfolding minus_diff_minus
+    apply (subst integral_diff(2)[symmetric])
+    apply (auto intro!: integral_cong_AE simp: split_beta' simp del: integral_diff)
+    done
 qed
 
 lemma (in information_space) conditional_entropy_eq:
-  assumes "simple_function M X" "simple_function M Z"
-  shows "\<H>(X | Z) =
-     - (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M.
-         joint_distribution X Z {(x, z)} *
-         log b (joint_distribution X Z {(x, z)} / distribution Z {z}))"
-  by (subst conditional_entropy_generic_eq[OF assms[THEN simple_function_imp_finite_random_variable]])
-     simp
+  assumes Y: "simple_distributed M Y Py" and X: "simple_function M X"
+  assumes XY: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
+    shows "\<H>(X | Y) = - (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))"
+proof (subst conditional_entropy_generic_eq[OF _ _
+  simple_distributed[OF simple_distributedI[OF X refl]] simple_distributed[OF Y] simple_distributed_joint[OF XY]])
+  have [simp]: "finite (X`space M)" using X by (simp add: simple_function_def)
+  note Y[THEN simple_distributed_finite, simp]
+  show "sigma_finite_measure (count_space (X ` space M))"
+    by (simp add: sigma_finite_measure_count_space_finite)
+  show "sigma_finite_measure (count_space (Y ` space M))"
+    by (simp add: sigma_finite_measure_count_space_finite)
+  let ?f = "(\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x else 0)"
+  have "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) = count_space (X`space M \<times> Y`space M)"
+    (is "?P = ?C")
+    using X Y by (simp add: simple_distributed_finite pair_measure_count_space)
+  with X Y show
+      "integrable ?P (\<lambda>x. ?f x * log b (?f x))"
+      "integrable ?P (\<lambda>x. ?f x * log b (Py (snd x)))"
+    by (auto intro!: integrable_count_space simp: simple_distributed_finite)
+  have eq: "(\<lambda>(x, y). ?f (x, y) * log b (?f (x, y) / Py y)) =
+    (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy x * log b (Pxy x / Py (snd x)) else 0)"
+    by auto
+  from X Y show "- (\<integral> (x, y). ?f (x, y) * log b (?f (x, y) / Py y) \<partial>?P) =
+    - (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` space M. Pxy (x, y) * log b (Pxy (x, y) / Py y))"
+    by (auto intro!: setsum_cong simp add: `?P = ?C` lebesgue_integral_count_space_finite simple_distributed_finite eq setsum_cases split_beta')
+qed
 
-lemma (in information_space) conditional_entropy_eq_ce_with_hypothesis:
+lemma (in information_space) conditional_mutual_information_eq_conditional_entropy:
   assumes X: "simple_function M X" and Y: "simple_function M Y"
-  shows "\<H>(X | Y) =
-    -(\<Sum>y\<in>Y`space M. distribution Y {y} *
-      (\<Sum>x\<in>X`space M. joint_distribution X Y {(x,y)} / distribution Y {(y)} *
-              log b (joint_distribution X Y {(x,y)} / distribution Y {(y)})))"
-  unfolding conditional_entropy_eq[OF assms]
-  using finite_distribution_order(5,6)[OF assms[THEN simple_function_imp_finite_random_variable]]
-  by (auto simp: setsum_cartesian_product'  setsum_commute[of _ "Y`space M"] setsum_right_distrib
-           intro!: setsum_cong)
+  shows "\<I>(X ; X | Y) = \<H>(X | Y)"
+proof -
+  def Py \<equiv> "\<lambda>x. if x \<in> Y`space M then measure M (Y -` {x} \<inter> space M) else 0"
+  def Pxy \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x))`space M then measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M) else 0"
+  def Pxxy \<equiv> "\<lambda>x. if x \<in> (\<lambda>x. (X x, X x, Y x))`space M then measure M ((\<lambda>x. (X x, X x, Y x)) -` {x} \<inter> space M) else 0"
+  let ?M = "X`space M \<times> X`space M \<times> Y`space M"
 
-lemma (in information_space) conditional_entropy_eq_cartesian_product:
-  assumes "simple_function M X" "simple_function M Y"
-  shows "\<H>(X | Y) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
-    joint_distribution X Y {(x,y)} *
-    log b (joint_distribution X Y {(x,y)} / distribution Y {y}))"
-  unfolding conditional_entropy_eq[OF assms]
-  by (auto intro!: setsum_cong simp: setsum_cartesian_product')
+  note XY = simple_function_Pair[OF X Y]
+  note XXY = simple_function_Pair[OF X XY]
+  have Py: "simple_distributed M Y Py"
+    using Y by (rule simple_distributedI) (auto simp: Py_def)
+  have Pxy: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
+    using XY by (rule simple_distributedI) (auto simp: Pxy_def)
+  have Pxxy: "simple_distributed M (\<lambda>x. (X x, X x, Y x)) Pxxy"
+    using XXY by (rule simple_distributedI) (auto simp: Pxxy_def)
+  have eq: "(\<lambda>x. (X x, X x, Y x)) ` space M = (\<lambda>(x, y). (x, x, y)) ` (\<lambda>x. (X x, Y x)) ` space M"
+    by auto
+  have inj: "\<And>A. inj_on (\<lambda>(x, y). (x, x, y)) A"
+    by (auto simp: inj_on_def)
+  have Pxxy_eq: "\<And>x y. Pxxy (x, x, y) = Pxy (x, y)"
+    by (auto simp: Pxxy_def Pxy_def intro!: arg_cong[where f=prob])
+  have "AE x in count_space ((\<lambda>x. (X x, Y x))`space M). Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+    by (intro subdensity_real[of snd, OF _ Pxy[THEN simple_distributed] Py[THEN simple_distributed]]) (auto intro: measurable_Pair)
+  then show ?thesis
+    apply (subst conditional_mutual_information_eq[OF Py Pxy Pxy Pxxy])
+    apply (subst conditional_entropy_eq[OF Py X Pxy])
+    apply (auto intro!: setsum_cong simp: Pxxy_eq setsum_negf[symmetric] eq setsum_reindex[OF inj]
+                log_simps zero_less_mult_iff zero_le_mult_iff field_simps mult_less_0_iff AE_count_space)
+    using Py[THEN simple_distributed, THEN distributed_real_AE] Pxy[THEN simple_distributed, THEN distributed_real_AE]
+    apply (auto simp add: not_le[symmetric] AE_count_space)
+    done
+qed
+
+lemma (in information_space) conditional_entropy_nonneg:
+  assumes X: "simple_function M X" and Y: "simple_function M Y" shows "0 \<le> \<H>(X | Y)"
+  using conditional_mutual_information_eq_conditional_entropy[OF X Y] conditional_mutual_information_nonneg[OF X X Y]
+  by simp
 
 subsection {* Equalities *}
 
-lemma (in information_space) mutual_information_eq_entropy_conditional_entropy:
-  assumes X: "simple_function M X" and Z: "simple_function M Z"
-  shows  "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)"
+lemma (in information_space) mutual_information_eq_entropy_conditional_entropy_distr:
+  fixes Px :: "'b \<Rightarrow> real" and Py :: "'c \<Rightarrow> real" and Pxy :: "('b \<times> 'c) \<Rightarrow> real"
+  assumes S: "sigma_finite_measure S" and T: "sigma_finite_measure T"
+  assumes Px: "distributed M S X Px" and Py: "distributed M T Y Py"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  assumes Ix: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Px (fst x)))"
+  assumes Iy: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Py (snd x)))"
+  assumes Ixy: "integrable(S \<Otimes>\<^isub>M T) (\<lambda>x. Pxy x * log b (Pxy x))"
+  shows  "mutual_information b S T X Y = entropy b S X + entropy b T Y - entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
 proof -
-  let ?XZ = "\<lambda>x z. joint_distribution X Z {(x, z)}"
-  let ?Z = "\<lambda>z. distribution Z {z}"
-  let ?X = "\<lambda>x. distribution X {x}"
-  note fX = X[THEN simple_function_imp_finite_random_variable]
-  note fZ = Z[THEN simple_function_imp_finite_random_variable]
-  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 zero_le_mult_iff field_simps less_le) }
-  note * = this
+  have X: "entropy b S X = - (\<integral>x. Pxy x * log b (Px (fst x)) \<partial>(S \<Otimes>\<^isub>M T))"
+    using b_gt_1 Px[THEN distributed_real_measurable]
+    apply (subst entropy_distr[OF S Px])
+    apply (subst distributed_transform_integral[OF Pxy Px, where T=fst])
+    apply (auto intro!: integral_cong)
+    done
+
+  have Y: "entropy b T Y = - (\<integral>x. Pxy x * log b (Py (snd x)) \<partial>(S \<Otimes>\<^isub>M T))"
+    using b_gt_1 Py[THEN distributed_real_measurable]
+    apply (subst entropy_distr[OF T Py])
+    apply (subst distributed_transform_integral[OF Pxy Py, where T=snd])
+    apply (auto intro!: integral_cong)
+    done
+
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T ..
+  have ST: "sigma_finite_measure (S \<Otimes>\<^isub>M T)" ..
+
+  have XY: "entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) = - (\<integral>x. Pxy x * log b (Pxy x) \<partial>(S \<Otimes>\<^isub>M T))"
+    by (subst entropy_distr[OF ST Pxy]) (auto intro!: integral_cong)
+  
+  have "AE x in S \<Otimes>\<^isub>M T. Px (fst x) = 0 \<longrightarrow> Pxy x = 0"
+    by (intro subdensity_real[of fst, OF _ Pxy Px]) (auto intro: measurable_Pair)
+  moreover have "AE x in S \<Otimes>\<^isub>M T. Py (snd x) = 0 \<longrightarrow> Pxy x = 0"
+    by (intro subdensity_real[of snd, OF _ Pxy Py]) (auto intro: measurable_Pair)
+  moreover have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Px (fst x)"
+    using Px by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_fst'' dest: distributed_real_AE distributed_real_measurable)
+  moreover have "AE x in S \<Otimes>\<^isub>M T. 0 \<le> Py (snd x)"
+    using Py by (intro ST.AE_pair_measure) (auto simp: comp_def intro!: measurable_snd'' dest: distributed_real_AE distributed_real_measurable)
+  moreover note Pxy[THEN distributed_real_AE]
+  ultimately have "AE x in S \<Otimes>\<^isub>M T. Pxy x * log b (Pxy x) - Pxy x * log b (Px (fst x)) - Pxy x * log b (Py (snd x)) = 
+    Pxy x * log b (Pxy x / (Px (fst x) * Py (snd x)))"
+    (is "AE x in _. ?f x = ?g x")
+  proof eventually_elim
+    case (goal1 x)
+    show ?case
+    proof cases
+      assume "Pxy x \<noteq> 0"
+      with goal1 have "0 < Px (fst x)" "0 < Py (snd x)" "0 < Pxy x"
+        by auto
+      then show ?thesis
+        using b_gt_1 by (simp add: log_simps mult_pos_pos less_imp_le field_simps)
+    qed simp
+  qed
+
+  have "entropy b S X + entropy b T Y - entropy b (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) = integral\<^isup>L (S \<Otimes>\<^isub>M T) ?f"
+    unfolding X Y XY
+    apply (subst integral_diff)
+    apply (intro integral_diff Ixy Ix Iy)+
+    apply (subst integral_diff)
+    apply (intro integral_diff Ixy Ix Iy)+
+    apply (simp add: field_simps)
+    done
+  also have "\<dots> = integral\<^isup>L (S \<Otimes>\<^isub>M T) ?g"
+    using `AE x in _. ?f x = ?g x` by (rule integral_cong_AE)
+  also have "\<dots> = mutual_information b S T X Y"
+    by (rule mutual_information_distr[OF S T Px Py Pxy, symmetric])
+  finally show ?thesis ..
+qed
+
+lemma (in information_space) mutual_information_eq_entropy_conditional_entropy:
+  assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y"
+  shows  "\<I>(X ; Y) = \<H>(X) - \<H>(X | Y)"
+proof -
+  have X: "simple_distributed M X (\<lambda>x. measure M (X -` {x} \<inter> space M))"
+    using sf_X by (rule simple_distributedI) auto
+  have Y: "simple_distributed M Y (\<lambda>x. measure M (Y -` {x} \<inter> space M))"
+    using sf_Y by (rule simple_distributedI) auto
+  have sf_XY: "simple_function M (\<lambda>x. (X x, Y x))"
+    using sf_X sf_Y by (rule simple_function_Pair)
+  then have XY: "simple_distributed M (\<lambda>x. (X x, Y x)) (\<lambda>x. measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M))"
+    by (rule simple_distributedI) auto
+  from simple_distributed_joint_finite[OF this, simp]
+  have eq: "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)"
+    by (simp add: pair_measure_count_space)
+
+  have "\<I>(X ; Y) = \<H>(X) + \<H>(Y) - entropy b (count_space (X`space M) \<Otimes>\<^isub>M count_space (Y`space M)) (\<lambda>x. (X x, Y x))"
+    using sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]
+    by (rule mutual_information_eq_entropy_conditional_entropy_distr) (auto simp: eq integrable_count_space)
+  then show ?thesis
+    unfolding conditional_entropy_def by simp
+qed
+
+lemma (in information_space) mutual_information_nonneg_simple:
+  assumes sf_X: "simple_function M X" and sf_Y: "simple_function M Y"
+  shows  "0 \<le> \<I>(X ; Y)"
+proof -
+  have X: "simple_distributed M X (\<lambda>x. measure M (X -` {x} \<inter> space M))"
+    using sf_X by (rule simple_distributedI) auto
+  have Y: "simple_distributed M Y (\<lambda>x. measure M (Y -` {x} \<inter> space M))"
+    using sf_Y by (rule simple_distributedI) auto
+
+  have sf_XY: "simple_function M (\<lambda>x. (X x, Y x))"
+    using sf_X sf_Y by (rule simple_function_Pair)
+  then have XY: "simple_distributed M (\<lambda>x. (X x, Y x)) (\<lambda>x. measure M ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M))"
+    by (rule simple_distributedI) auto
+
+  from simple_distributed_joint_finite[OF this, simp]
+  have eq: "count_space (X ` space M) \<Otimes>\<^isub>M count_space (Y ` space M) = count_space (X ` space M \<times> Y ` space M)"
+    by (simp add: pair_measure_count_space)
+
   show ?thesis
-    unfolding entropy_eq[OF X] conditional_entropy_eq[OF X Z] mutual_information_eq[OF X Z]
-    using setsum_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_distribution)
+    by (rule mutual_information_nonneg[OF _ _ simple_distributed[OF X] simple_distributed[OF Y] simple_distributed_joint[OF XY]])
+       (simp_all add: eq integrable_count_space sigma_finite_measure_count_space_finite)
 qed
 
 lemma (in information_space) conditional_entropy_less_eq_entropy:
   assumes X: "simple_function M X" and Z: "simple_function M 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[OF assms] .
-  with mutual_information_positive[OF X Z] entropy_positive[OF X]
-  show ?thesis by auto
+  have "0 \<le> \<I>(X ; Z)" using X Z by (rule mutual_information_nonneg_simple)
+  also have "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)" using mutual_information_eq_entropy_conditional_entropy[OF assms] .
+  finally show ?thesis by auto
 qed
 
 lemma (in information_space) entropy_chain_rule:
   assumes X: "simple_function M X" and Y: "simple_function M Y"
   shows  "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)"
 proof -
-  let ?XY = "\<lambda>x y. joint_distribution X Y {(x, y)}"
-  let ?Y = "\<lambda>y. distribution Y {y}"
-  let ?X = "\<lambda>x. distribution X {x}"
-  note fX = X[THEN simple_function_imp_finite_random_variable]
-  note fY = Y[THEN simple_function_imp_finite_random_variable]
-  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 zero_le_mult_iff field_simps less_le) }
-  note * = this
-  show ?thesis
-    using setsum_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 *}
-
-definition "subvimage A f g \<longleftrightarrow> (\<forall>x \<in> A. f -` {f x} \<inter> A \<subseteq> g -` {g x} \<inter> A)"
-
-lemma subvimageI:
-  assumes "\<And>x y. \<lbrakk> x \<in> A ; y \<in> A ; f x = f y \<rbrakk> \<Longrightarrow> g x = g y"
-  shows "subvimage A f g"
-  using assms unfolding subvimage_def by blast
-
-lemma subvimageE[consumes 1]:
-  assumes "subvimage A f g"
-  obtains "\<And>x y. \<lbrakk> x \<in> A ; y \<in> A ; f x = f y \<rbrakk> \<Longrightarrow> g x = g y"
-  using assms unfolding subvimage_def by blast
-
-lemma subvimageD:
-  "\<lbrakk> subvimage A f g ; x \<in> A ; y \<in> A ; f x = f y \<rbrakk> \<Longrightarrow> g x = g y"
-  using assms unfolding subvimage_def by blast
-
-lemma subvimage_subset:
-  "\<lbrakk> subvimage B f g ; A \<subseteq> B \<rbrakk> \<Longrightarrow> subvimage A f g"
-  unfolding subvimage_def by auto
-
-lemma subvimage_idem[intro]: "subvimage A g g"
-  by (safe intro!: subvimageI)
-
-lemma subvimage_comp_finer[intro]:
-  assumes svi: "subvimage A g h"
-  shows "subvimage A g (f \<circ> h)"
-proof (rule subvimageI, simp)
-  fix x y assume "x \<in> A" "y \<in> A" "g x = g y"
-  from svi[THEN subvimageD, OF this]
-  show "f (h x) = f (h y)" by simp
-qed
-
-lemma subvimage_comp_gran:
-  assumes svi: "subvimage A g h"
-  assumes inj: "inj_on f (g ` A)"
-  shows "subvimage A (f \<circ> g) h"
-  by (rule subvimageI) (auto intro!: subvimageD[OF svi] simp: inj_on_iff[OF inj])
-
-lemma subvimage_comp:
-  assumes svi: "subvimage (f ` A) g h"
-  shows "subvimage A (g \<circ> f) (h \<circ> f)"
-  by (rule subvimageI) (auto intro!: svi[THEN subvimageD])
-
-lemma subvimage_trans:
-  assumes fg: "subvimage A f g"
-  assumes gh: "subvimage A g h"
-  shows "subvimage A f h"
-  by (rule subvimageI) (auto intro!: fg[THEN subvimageD] gh[THEN subvimageD])
-
-lemma subvimage_translator:
-  assumes svi: "subvimage A f g"
-  shows "\<exists>h. \<forall>x \<in> A. h (f x)  = g x"
-proof (safe intro!: exI[of _ "\<lambda>x. (THE z. z \<in> (g ` (f -` {x} \<inter> A)))"])
-  fix x assume "x \<in> A"
-  show "(THE x'. x' \<in> (g ` (f -` {f x} \<inter> A))) = g x"
-    by (rule theI2[of _ "g x"])
-      (insert `x \<in> A`, auto intro!: svi[THEN subvimageD])
-qed
-
-lemma subvimage_translator_image:
-  assumes svi: "subvimage A f g"
-  shows "\<exists>h. h ` f ` A = g ` A"
-proof -
-  from subvimage_translator[OF svi]
-  obtain h where "\<And>x. x \<in> A \<Longrightarrow> h (f x) = g x" by auto
-  thus ?thesis
-    by (auto intro!: exI[of _ h]
-      simp: image_compose[symmetric] comp_def cong: image_cong)
-qed
-
-lemma subvimage_finite:
-  assumes svi: "subvimage A f g" and fin: "finite (f`A)"
-  shows "finite (g`A)"
-proof -
-  from subvimage_translator_image[OF svi]
-  obtain h where "g`A = h`f`A" by fastforce
-  with fin show "finite (g`A)" by simp
-qed
-
-lemma subvimage_disj:
-  assumes svi: "subvimage A f g"
-  shows "f -` {x} \<inter> A \<subseteq> g -` {y} \<inter> A \<or>
-      f -` {x} \<inter> g -` {y} \<inter> A = {}" (is "?sub \<or> ?dist")
-proof (rule disjCI)
-  assume "\<not> ?dist"
-  then obtain z where "z \<in> A" and "x = f z" and "y = g z" by auto
-  thus "?sub" using svi unfolding subvimage_def by auto
-qed
-
-lemma setsum_image_split:
-  assumes svi: "subvimage A f g" and fin: "finite (f ` A)"
-  shows "(\<Sum>x\<in>f`A. h x) = (\<Sum>y\<in>g`A. \<Sum>x\<in>f`(g -` {y} \<inter> A). h x)"
-    (is "?lhs = ?rhs")
-proof -
-  have "f ` A =
-      snd ` (SIGMA x : g ` A. f ` (g -` {x} \<inter> A))"
-      (is "_ = snd ` ?SIGMA")
-    unfolding image_split_eq_Sigma[symmetric]
-    by (simp add: image_compose[symmetric] comp_def)
-  moreover
-  have snd_inj: "inj_on snd ?SIGMA"
-    unfolding image_split_eq_Sigma[symmetric]
-    by (auto intro!: inj_onI subvimageD[OF svi])
-  ultimately
-  have "(\<Sum>x\<in>f`A. h x) = (\<Sum>(x,y)\<in>?SIGMA. h y)"
-    by (auto simp: setsum_reindex intro: setsum_cong)
-  also have "... = ?rhs"
-    using subvimage_finite[OF svi fin] fin
-    apply (subst setsum_Sigma[symmetric])
-    by (auto intro!: finite_subset[of _ "f`A"])
-  finally show ?thesis .
+  note XY = simple_distributedI[OF simple_function_Pair[OF X Y] refl]
+  note YX = simple_distributedI[OF simple_function_Pair[OF Y X] refl]
+  note simple_distributed_joint_finite[OF this, simp]
+  let ?f = "\<lambda>x. prob ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M)"
+  let ?g = "\<lambda>x. prob ((\<lambda>x. (Y x, X x)) -` {x} \<inter> space M)"
+  let ?h = "\<lambda>x. if x \<in> (\<lambda>x. (Y x, X x)) ` space M then prob ((\<lambda>x. (Y x, X x)) -` {x} \<inter> space M) else 0"
+  have "\<H>(\<lambda>x. (X x, Y x)) = - (\<Sum>x\<in>(\<lambda>x. (X x, Y x)) ` space M. ?f x * log b (?f x))"
+    using XY by (rule entropy_simple_distributed)
+  also have "\<dots> = - (\<Sum>x\<in>(\<lambda>(x, y). (y, x)) ` (\<lambda>x. (X x, Y x)) ` space M. ?g x * log b (?g x))"
+    by (subst (2) setsum_reindex) (auto simp: inj_on_def intro!: setsum_cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
+  also have "\<dots> = - (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` space M. ?h x * log b (?h x))"
+    by (auto intro!: setsum_cong)
+  also have "\<dots> = entropy b (count_space (Y ` space M) \<Otimes>\<^isub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))"
+    by (subst entropy_distr[OF _ simple_distributed_joint[OF YX]])
+       (auto simp: pair_measure_count_space sigma_finite_measure_count_space_finite lebesgue_integral_count_space_finite
+             cong del: setsum_cong  intro!: setsum_mono_zero_left)
+  finally have "\<H>(\<lambda>x. (X x, Y x)) = entropy b (count_space (Y ` space M) \<Otimes>\<^isub>M count_space (X ` space M)) (\<lambda>x. (Y x, X x))" .
+  then show ?thesis
+    unfolding conditional_entropy_def by simp
 qed
 
 lemma (in information_space) entropy_partition:
-  assumes sf: "simple_function M X" "simple_function M P"
-  assumes svi: "subvimage (space M) X P"
-  shows "\<H>(X) = \<H>(P) + \<H>(X|P)"
+  assumes X: "simple_function M X"
+  shows "\<H>(X) = \<H>(f \<circ> X) + \<H>(X|f \<circ> X)"
 proof -
-  let ?XP = "\<lambda>x p. joint_distribution X P {(x, p)}"
-  let ?X = "\<lambda>x. distribution X {x}"
-  let ?P = "\<lambda>p. 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 finite_distribution_order[OF fX fP, simp]
-  have "(\<Sum>x\<in>X ` space M. ?X x * log b (?X x)) =
-    (\<Sum>y\<in>P `space M. \<Sum>x\<in>X ` space M. ?XP x y * log b (?XP x y))"
-  proof (subst setsum_image_split[OF svi],
-      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 "?XP (X x) (P p) * log b (?XP (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)
-    with svi[unfolded subvimage_def, rule_format, OF `x \<in> space M`]
-    show "x \<in> P -` {P p}" by auto
-  next
-    fix p x assume in_space: "p \<in> space M" "x \<in> space M"
-    assume "P x = P p"
-    from this[symmetric] svi[unfolded subvimage_def, rule_format, OF `x \<in> space M`]
-    have "X -` {X x} \<inter> space M \<subseteq> P -` {P p} \<inter> space M"
-      by auto
-    hence "(\<lambda>x. (X x, P x)) -` {(X x, P p)} \<inter> space M = X -` {X x} \<inter> space M"
-      by auto
-    thus "?X (X x) * log b (?X (X x)) = ?XP (X x) (P p) * log b (?XP (X x) (P p))"
-      by (auto simp: distribution_def)
-  qed
-  moreover have "\<And>x y. ?XP x y * log b (?XP x y / ?P y) =
-      ?XP x y * log b (?XP x y) - ?XP x y * log b (?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_joint_distribution_singleton[OF fX fP]
-    by (simp add: setsum_cartesian_product' setsum_subtractf setsum_distribution
-      setsum_left_distrib[symmetric] setsum_commute[where B="P`space M"])
+  note fX = simple_function_compose[OF X, of f]  
+  have eq: "(\<lambda>x. ((f \<circ> X) x, X x)) ` space M = (\<lambda>x. (f x, x)) ` X ` space M" by auto
+  have inj: "\<And>A. inj_on (\<lambda>x. (f x, x)) A"
+    by (auto simp: inj_on_def)
+  show ?thesis
+    apply (subst entropy_chain_rule[symmetric, OF fX X])
+    apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_Pair[OF fX X] refl]])
+    apply (subst entropy_simple_distributed[OF simple_distributedI[OF X refl]])
+    unfolding eq
+    apply (subst setsum_reindex[OF inj])
+    apply (auto intro!: setsum_cong arg_cong[where f="\<lambda>A. prob A * log b (prob A)"])
+    done
 qed
 
 corollary (in information_space) entropy_data_processing:
   assumes X: "simple_function M X" shows "\<H>(f \<circ> X) \<le> \<H>(X)"
 proof -
-  note X
-  moreover have fX: "simple_function M (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)
+  note fX = simple_function_compose[OF X, of f]
+  from X 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])
+    by (auto intro: conditional_entropy_nonneg[OF X fX])
 qed
 
 corollary (in information_space) entropy_of_inj:
@@ -1411,7 +1313,11 @@
   have sf: "simple_function M (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 inj])
+    unfolding o_assoc
+    apply (subst entropy_simple_distributed[OF simple_distributedI[OF X refl]])
+    apply (subst entropy_simple_distributed[OF simple_distributedI[OF simple_function_compose[OF X]], where f="\<lambda>x. prob (X -` {x} \<inter> space M)"])
+    apply (auto intro!: setsum_cong arg_cong[where f=prob] image_eqI simp: the_inv_into_f_f[OF inj] comp_def)
+    done
   also have "... \<le> \<H>(f \<circ> X)"
     using entropy_data_processing[OF sf] .
   finally show "\<H>(X) \<le> \<H>(f \<circ> X)" .
--- a/src/HOL/Probability/Lebesgue_Integration.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Lebesgue_Integration.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -6,9 +6,13 @@
 header {*Lebesgue Integration*}
 
 theory Lebesgue_Integration
-  imports Measure Borel_Space
+  imports Measure_Space Borel_Space
 begin
 
+lemma ereal_minus_eq_PInfty_iff:
+  fixes x y :: ereal shows "x - y = \<infinity> \<longleftrightarrow> y = -\<infinity> \<or> x = \<infinity>"
+  by (cases x y rule: ereal2_cases) simp_all
+
 lemma real_ereal_1[simp]: "real (1::ereal) = 1"
   unfolding one_ereal_def by simp
 
@@ -28,17 +32,17 @@
     by (intro tendsto_add assms tendsto_divide tendsto_norm tendsto_diff) auto
 qed
 
-lemma (in measure_space) measure_Union:
+lemma measure_Union:
   assumes "finite S" "S \<subseteq> sets M" "\<And>A B. A \<in> S \<Longrightarrow> B \<in> S \<Longrightarrow> A \<noteq> B \<Longrightarrow> A \<inter> B = {}"
-  shows "setsum \<mu> S = \<mu> (\<Union>S)"
+  shows "setsum (emeasure M) S = (emeasure M) (\<Union>S)"
 proof -
-  have "setsum \<mu> S = \<mu> (\<Union>i\<in>S. i)"
-    using assms by (intro measure_setsum[OF `finite S`]) (auto simp: disjoint_family_on_def)
-  also have "\<dots> = \<mu> (\<Union>S)" by (auto intro!: arg_cong[where f=\<mu>])
+  have "setsum (emeasure M) S = (emeasure M) (\<Union>i\<in>S. i)"
+    using assms by (intro setsum_emeasure[OF _ _ `finite S`]) (auto simp: disjoint_family_on_def)
+  also have "\<dots> = (emeasure M) (\<Union>S)" by (auto intro!: arg_cong[where f="emeasure M"])
   finally show ?thesis .
 qed
 
-lemma (in sigma_algebra) measurable_sets2[intro]:
+lemma measurable_sets2[intro]:
   assumes "f \<in> measurable M M'" "g \<in> measurable M M''"
   and "A \<in> sets M'" "B \<in> sets M''"
   shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
@@ -55,7 +59,7 @@
 
 lemma borel_measurable_real_floor:
   "(\<lambda>x::real. real \<lfloor>x\<rfloor>) \<in> borel_measurable borel"
-  unfolding borel.borel_measurable_iff_ge
+  unfolding borel_measurable_iff_ge
 proof (intro allI)
   fix a :: real
   { fix x have "a \<le> real \<lfloor>x\<rfloor> \<longleftrightarrow> real \<lceil>a\<rceil> \<le> x"
@@ -65,19 +69,7 @@
   then show "{w::real \<in> space borel. a \<le> real \<lfloor>w\<rfloor>} \<in> sets borel" by auto
 qed
 
-lemma measure_preservingD2:
-  "f \<in> measure_preserving A B \<Longrightarrow> f \<in> measurable A B"
-  unfolding measure_preserving_def by auto
-
-lemma measure_preservingD3:
-  "f \<in> measure_preserving A B \<Longrightarrow> f \<in> space A \<rightarrow> space B"
-  unfolding measure_preserving_def measurable_def by auto
-
-lemma measure_preservingD:
-  "T \<in> measure_preserving A B \<Longrightarrow> X \<in> sets B \<Longrightarrow> measure A (T -` X \<inter> space A) = measure B X"
-  unfolding measure_preserving_def by auto
-
-lemma (in sigma_algebra) borel_measurable_real_natfloor[intro, simp]:
+lemma borel_measurable_real_natfloor[intro, simp]:
   assumes "f \<in> borel_measurable M"
   shows "(\<lambda>x. real (natfloor (f x))) \<in> borel_measurable M"
 proof -
@@ -87,8 +79,8 @@
   show ?thesis by (simp add: comp_def)
 qed
 
-lemma (in measure_space) AE_not_in:
-  assumes N: "N \<in> null_sets" shows "AE x. x \<notin> N"
+lemma AE_not_in:
+  assumes N: "N \<in> null_sets M" shows "AE x in M. x \<notin> N"
   using N by (rule AE_I') auto
 
 lemma sums_If_finite:
@@ -128,7 +120,7 @@
     finite (g ` space M) \<and>
     (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
 
-lemma (in sigma_algebra) simple_functionD:
+lemma simple_functionD:
   assumes "simple_function M g"
   shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M"
 proof -
@@ -140,7 +132,7 @@
     by (auto intro!: finite_UN simp del: UN_simps simp: simple_function_def)
 qed
 
-lemma (in sigma_algebra) simple_function_measurable2[intro]:
+lemma simple_function_measurable2[intro]:
   assumes "simple_function M f" "simple_function M g"
   shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M"
 proof -
@@ -149,7 +141,7 @@
   then show ?thesis using assms[THEN simple_functionD(2)] by auto
 qed
 
-lemma (in sigma_algebra) simple_function_indicator_representation:
+lemma simple_function_indicator_representation:
   fixes f ::"'a \<Rightarrow> ereal"
   assumes f: "simple_function M f" and x: "x \<in> space M"
   shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
@@ -164,7 +156,7 @@
   finally show ?thesis by auto
 qed
 
-lemma (in measure_space) simple_function_notspace:
+lemma simple_function_notspace:
   "simple_function M (\<lambda>x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h")
 proof -
   have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
@@ -173,7 +165,7 @@
   thus ?thesis unfolding simple_function_def by auto
 qed
 
-lemma (in sigma_algebra) simple_function_cong:
+lemma simple_function_cong:
   assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
   shows "simple_function M f \<longleftrightarrow> simple_function M g"
 proof -
@@ -183,12 +175,12 @@
   thus ?thesis unfolding simple_function_def using assms by simp
 qed
 
-lemma (in sigma_algebra) simple_function_cong_algebra:
+lemma simple_function_cong_algebra:
   assumes "sets N = sets M" "space N = space M"
   shows "simple_function M f \<longleftrightarrow> simple_function N f"
   unfolding simple_function_def assms ..
 
-lemma (in sigma_algebra) borel_measurable_simple_function:
+lemma borel_measurable_simple_function:
   assumes "simple_function M f"
   shows "f \<in> borel_measurable M"
 proof (rule borel_measurableI)
@@ -204,24 +196,23 @@
   thus "f -` S \<inter> space M \<in> sets M" unfolding * .
 qed
 
-lemma (in sigma_algebra) simple_function_borel_measurable:
+lemma simple_function_borel_measurable:
   fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
   assumes "f \<in> borel_measurable M" and "finite (f ` space M)"
   shows "simple_function M f"
   using assms unfolding simple_function_def
   by (auto intro: borel_measurable_vimage)
 
-lemma (in sigma_algebra) simple_function_eq_borel_measurable:
+lemma simple_function_eq_borel_measurable:
   fixes f :: "'a \<Rightarrow> ereal"
   shows "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> borel_measurable M"
-  using simple_function_borel_measurable[of f]
-    borel_measurable_simple_function[of f]
+  using simple_function_borel_measurable[of f] borel_measurable_simple_function[of M f]
   by (fastforce simp: simple_function_def)
 
-lemma (in sigma_algebra) simple_function_const[intro, simp]:
+lemma simple_function_const[intro, simp]:
   "simple_function M (\<lambda>x. c)"
   by (auto intro: finite_subset simp: simple_function_def)
-lemma (in sigma_algebra) simple_function_compose[intro, simp]:
+lemma simple_function_compose[intro, simp]:
   assumes "simple_function M f"
   shows "simple_function M (g \<circ> f)"
   unfolding simple_function_def
@@ -238,7 +229,7 @@
     by (rule_tac finite_UN) (auto intro!: finite_UN)
 qed
 
-lemma (in sigma_algebra) simple_function_indicator[intro, simp]:
+lemma simple_function_indicator[intro, simp]:
   assumes "A \<in> sets M"
   shows "simple_function M (indicator A)"
 proof -
@@ -250,7 +241,7 @@
     using assms by (auto simp: indicator_def [abs_def])
 qed
 
-lemma (in sigma_algebra) simple_function_Pair[intro, simp]:
+lemma simple_function_Pair[intro, simp]:
   assumes "simple_function M f"
   assumes "simple_function M g"
   shows "simple_function M (\<lambda>x. (f x, g x))" (is "simple_function M ?p")
@@ -268,13 +259,13 @@
     using assms unfolding simple_function_def by auto
 qed
 
-lemma (in sigma_algebra) simple_function_compose1:
+lemma simple_function_compose1:
   assumes "simple_function M f"
   shows "simple_function M (\<lambda>x. g (f x))"
   using simple_function_compose[OF assms, of g]
   by (simp add: comp_def)
 
-lemma (in sigma_algebra) simple_function_compose2:
+lemma simple_function_compose2:
   assumes "simple_function M f" and "simple_function M g"
   shows "simple_function M (\<lambda>x. h (f x) (g x))"
 proof -
@@ -283,7 +274,7 @@
   thus ?thesis by (simp_all add: comp_def)
 qed
 
-lemmas (in sigma_algebra) simple_function_add[intro, simp] = simple_function_compose2[where h="op +"]
+lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"]
   and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"]
   and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"]
   and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"]
@@ -291,24 +282,24 @@
   and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"]
   and simple_function_max[intro, simp] = simple_function_compose2[where h=max]
 
-lemma (in sigma_algebra) simple_function_setsum[intro, simp]:
+lemma simple_function_setsum[intro, simp]:
   assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
   shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)"
 proof cases
   assume "finite P" from this assms show ?thesis by induct auto
 qed auto
 
-lemma (in sigma_algebra)
+lemma
   fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f"
   shows simple_function_ereal[intro, simp]: "simple_function M (\<lambda>x. ereal (f x))"
   by (auto intro!: simple_function_compose1[OF sf])
 
-lemma (in sigma_algebra)
+lemma
   fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f"
   shows simple_function_real_of_nat[intro, simp]: "simple_function M (\<lambda>x. real (f x))"
   by (auto intro!: simple_function_compose1[OF sf])
 
-lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
+lemma borel_measurable_implies_simple_function_sequence:
   fixes u :: "'a \<Rightarrow> ereal"
   assumes u: "u \<in> borel_measurable M"
   shows "\<exists>f. incseq f \<and> (\<forall>i. \<infinity> \<notin> range (f i) \<and> simple_function M (f i)) \<and>
@@ -416,14 +407,14 @@
   qed (auto simp: divide_nonneg_pos)
 qed
 
-lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence':
+lemma borel_measurable_implies_simple_function_sequence':
   fixes u :: "'a \<Rightarrow> ereal"
   assumes u: "u \<in> borel_measurable M"
   obtains f where "\<And>i. simple_function M (f i)" "incseq f" "\<And>i. \<infinity> \<notin> range (f i)"
     "\<And>x. (SUP i. f i x) = max 0 (u x)" "\<And>i x. 0 \<le> f i x"
   using borel_measurable_implies_simple_function_sequence[OF u] by auto
 
-lemma (in sigma_algebra) simple_function_If_set:
+lemma simple_function_If_set:
   assumes sf: "simple_function M f" "simple_function M g" and A: "A \<inter> space M \<in> sets M"
   shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF")
 proof -
@@ -445,7 +436,7 @@
   qed
 qed
 
-lemma (in sigma_algebra) simple_function_If:
+lemma simple_function_If:
   assumes sf: "simple_function M f" "simple_function M g" and P: "{x\<in>space M. P x} \<in> sets M"
   shows "simple_function M (\<lambda>x. if P x then f x else g x)"
 proof -
@@ -453,58 +444,17 @@
   with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
 qed
 
-lemma (in measure_space) simple_function_restricted:
-  fixes f :: "'a \<Rightarrow> ereal" assumes "A \<in> sets M"
-  shows "simple_function (restricted_space A) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator A x)"
-    (is "simple_function ?R f \<longleftrightarrow> simple_function M ?f")
-proof -
-  interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
-  have f: "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
-  proof cases
-    assume "A = space M"
-    then have "f`A = ?f`space M" by (fastforce simp: image_iff)
-    then show ?thesis by simp
-  next
-    assume "A \<noteq> space M"
-    then obtain x where x: "x \<in> space M" "x \<notin> A"
-      using sets_into_space `A \<in> sets M` by auto
-    have *: "?f`space M = f`A \<union> {0}"
-    proof (auto simp add: image_iff)
-      show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0"
-        using x by (auto intro!: bexI[of _ x])
-    next
-      fix x assume "x \<in> A"
-      then show "\<exists>y\<in>space M. f x = f y * indicator A y"
-        using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
-    next
-      fix x
-      assume "indicator A x \<noteq> (0::ereal)"
-      then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
-      moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
-      ultimately show "f x = 0" by auto
-    qed
-    then show ?thesis by auto
-  qed
-  then show ?thesis
-    unfolding simple_function_eq_borel_measurable
-      R.simple_function_eq_borel_measurable
-    unfolding borel_measurable_restricted[OF `A \<in> sets M`]
-    using assms(1)[THEN sets_into_space]
-    by (auto simp: indicator_def)
-qed
-
-lemma (in sigma_algebra) simple_function_subalgebra:
+lemma simple_function_subalgebra:
   assumes "simple_function N f"
   and N_subalgebra: "sets N \<subseteq> sets M" "space N = space M"
   shows "simple_function M f"
   using assms unfolding simple_function_def by auto
 
-lemma (in measure_space) simple_function_vimage:
-  assumes T: "sigma_algebra M'" "T \<in> measurable M M'"
+lemma simple_function_comp:
+  assumes T: "T \<in> measurable M M'"
     and f: "simple_function M' f"
   shows "simple_function M (\<lambda>x. f (T x))"
 proof (intro simple_function_def[THEN iffD2] conjI ballI)
-  interpret T: sigma_algebra M' by fact
   have "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
     using T unfolding measurable_def by auto
   then show "finite ((\<lambda>x. f (T x)) ` space M)"
@@ -523,16 +473,16 @@
 
 section "Simple integral"
 
-definition simple_integral_def:
-  "integral\<^isup>S M f = (\<Sum>x \<in> f ` space M. x * measure M (f -` {x} \<inter> space M))"
+definition simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^isup>S") where
+  "integral\<^isup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
 
 syntax
-  "_simple_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ereal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
+  "_simple_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^isup>S _. _ \<partial>_" [60,61] 110)
 
 translations
-  "\<integral>\<^isup>S x. f \<partial>M" == "CONST integral\<^isup>S M (%x. f)"
+  "\<integral>\<^isup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)"
 
-lemma (in measure_space) simple_integral_cong:
+lemma simple_integral_cong:
   assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"
   shows "integral\<^isup>S M f = integral\<^isup>S M g"
 proof -
@@ -542,19 +492,8 @@
   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> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
-    and "simple_function M f"
-  shows "integral\<^isup>S N f = integral\<^isup>S M f"
-proof -
-  interpret v: measure_space N
-    by (rule measure_space_cong) fact+
-  from simple_functionD[OF `simple_function M f`] assms show ?thesis
-    by (auto intro!: setsum_cong simp: simple_integral_def)
-qed
-
-lemma (in measure_space) simple_integral_const[simp]:
-  "(\<integral>\<^isup>Sx. c \<partial>M) = c * \<mu> (space M)"
+lemma simple_integral_const[simp]:
+  "(\<integral>\<^isup>Sx. c \<partial>M) = c * (emeasure M) (space M)"
 proof (cases "space M = {}")
   case True thus ?thesis unfolding simple_integral_def by simp
 next
@@ -562,9 +501,9 @@
   thus ?thesis unfolding simple_integral_def by simp
 qed
 
-lemma (in measure_space) simple_function_partition:
+lemma simple_function_partition:
   assumes f: "simple_function M f" and g: "simple_function M g"
-  shows "integral\<^isup>S M f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. the_elem (f`A) * \<mu> A)"
+  shows "integral\<^isup>S M f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. the_elem (f`A) * (emeasure M) A)"
     (is "_ = setsum _ (?p ` space M)")
 proof-
   let ?sub = "\<lambda>x. ?p ` (f -` {x} \<inter> space M)"
@@ -586,13 +525,13 @@
 
   { fix x assume "x \<in> space M"
     have "\<Union>(?sub (f x)) = (f -` {f x} \<inter> space M)" by auto
-    with sets have "\<mu> (f -` {f x} \<inter> space M) = setsum \<mu> (?sub (f x))"
+    with sets have "(emeasure M) (f -` {f x} \<inter> space M) = setsum (emeasure M) (?sub (f x))"
       by (subst measure_Union) auto }
-  hence "integral\<^isup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)"
+  hence "integral\<^isup>S M f = (\<Sum>(x,A)\<in>?SIGMA. x * (emeasure M) A)"
     unfolding simple_integral_def using f sets
     by (subst setsum_Sigma[symmetric])
        (auto intro!: setsum_cong setsum_ereal_right_distrib)
-  also have "\<dots> = (\<Sum>A\<in>?p ` space M. the_elem (f`A) * \<mu> A)"
+  also have "\<dots> = (\<Sum>A\<in>?p ` space M. the_elem (f`A) * (emeasure M) A)"
   proof -
     have [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f ` ?p x = {f x}" by (auto intro!: imageI)
     have "(\<lambda>A. (the_elem (f ` A), A)) ` ?p ` space M
@@ -610,7 +549,7 @@
   finally show ?thesis .
 qed
 
-lemma (in measure_space) simple_integral_add[simp]:
+lemma simple_integral_add[simp]:
   assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<le> g x"
   shows "(\<integral>\<^isup>Sx. f x + g x \<partial>M) = integral\<^isup>S M f + integral\<^isup>S M g"
 proof -
@@ -628,7 +567,7 @@
        (auto simp add: ereal_left_distrib setsum_addf[symmetric] intro!: setsum_cong)
 qed
 
-lemma (in measure_space) simple_integral_setsum[simp]:
+lemma simple_integral_setsum[simp]:
   assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x"
   assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)"
   shows "(\<integral>\<^isup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>S M (f i))"
@@ -638,11 +577,11 @@
     by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg)
 qed auto
 
-lemma (in measure_space) simple_integral_mult[simp]:
+lemma simple_integral_mult[simp]:
   assumes f: "simple_function M f" "\<And>x. 0 \<le> f x"
   shows "(\<integral>\<^isup>Sx. c * f x \<partial>M) = c * integral\<^isup>S M f"
 proof -
-  note mult = simple_function_mult[OF simple_function_const[of c] f(1)]
+  note mult = simple_function_mult[OF simple_function_const[of _ c] f(1)]
   { fix x let ?S = "f -` {f x} \<inter> (\<lambda>x. c * f x) -` {c * f x} \<inter> space M"
     assume "x \<in> space M"
     hence "(\<lambda>x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}"
@@ -654,9 +593,9 @@
        (auto intro!: ereal_0_le_mult setsum_cong simp: mult_assoc)
 qed
 
-lemma (in measure_space) simple_integral_mono_AE:
+lemma simple_integral_mono_AE:
   assumes f: "simple_function M f" and g: "simple_function M g"
-  and mono: "AE x. f x \<le> g x"
+  and mono: "AE x in M. f x \<le> g x"
   shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
 proof -
   let ?S = "\<lambda>x. (g -` {g x} \<inter> space M) \<inter> (f -` {f x} \<inter> space M)"
@@ -669,55 +608,55 @@
   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)"
+    show "the_elem (f`?S x) * (emeasure M) (?S x) \<le> the_elem (g`?S x) * (emeasure M) (?S x)"
     proof (cases "f x \<le> g x")
       case True then show ?thesis
         using * assms(1,2)[THEN simple_functionD(2)]
         by (auto intro!: ereal_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"
+      obtain N where N: "{x\<in>space M. \<not> f x \<le> g x} \<subseteq> N" "N \<in> sets M" "(emeasure M) 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
         by (rule_tac Int) (auto intro!: simple_functionD)
-      ultimately have "\<mu> (?S x) \<le> \<mu> N"
-        using `N \<in> sets M` by (auto intro!: measure_mono)
-      moreover have "0 \<le> \<mu> (?S x)"
+      ultimately have "(emeasure M) (?S x) \<le> (emeasure M) N"
+        using `N \<in> sets M` by (auto intro!: emeasure_mono)
+      moreover have "0 \<le> (emeasure M) (?S x)"
         using assms(1,2)[THEN simple_functionD(2)] by auto
-      ultimately have "\<mu> (?S x) = 0" using `\<mu> N = 0` by auto
+      ultimately have "(emeasure M) (?S x) = 0" using `(emeasure M) N = 0` by auto
       then show ?thesis by simp
     qed
   qed
 qed
 
-lemma (in measure_space) simple_integral_mono:
+lemma simple_integral_mono:
   assumes "simple_function M f" and "simple_function M g"
   and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"
   shows "integral\<^isup>S M f \<le> integral\<^isup>S M g"
   using assms by (intro simple_integral_mono_AE) auto
 
-lemma (in measure_space) simple_integral_cong_AE:
+lemma simple_integral_cong_AE:
   assumes "simple_function M f" and "simple_function M g"
-  and "AE x. f x = g x"
+  and "AE x in M. f x = g x"
   shows "integral\<^isup>S M f = integral\<^isup>S M g"
   using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE)
 
-lemma (in measure_space) simple_integral_cong':
+lemma simple_integral_cong':
   assumes sf: "simple_function M f" "simple_function M g"
-  and mea: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0"
+  and mea: "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0"
   shows "integral\<^isup>S M f = integral\<^isup>S M 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 "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0" by fact
   show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
     using sf[THEN borel_measurable_simple_function] by auto
 qed simp
 
-lemma (in measure_space) simple_integral_indicator:
+lemma simple_integral_indicator:
   assumes "A \<in> sets M"
   assumes "simple_function M f"
   shows "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
-    (\<Sum>x \<in> f ` space M. x * \<mu> (f -` {x} \<inter> space M \<inter> A))"
+    (\<Sum>x \<in> f ` space M. x * (emeasure M) (f -` {x} \<inter> space M \<inter> A))"
 proof cases
   assume "A = space M"
   moreover hence "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) = integral\<^isup>S M f"
@@ -737,7 +676,7 @@
     show "0 \<in> ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x])
   qed
   have *: "(\<integral>\<^isup>Sx. f x * indicator A x \<partial>M) =
-    (\<Sum>x \<in> f ` space M \<union> {0}. x * \<mu> (f -` {x} \<inter> space M \<inter> A))"
+    (\<Sum>x \<in> f ` space M \<union> {0}. x * (emeasure M) (f -` {x} \<inter> space M \<inter> A))"
     unfolding simple_integral_def I
   proof (rule setsum_mono_zero_cong_left)
     show "finite (f ` space M \<union> {0})"
@@ -747,118 +686,83 @@
     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
+      i * (emeasure M) (f -` {i} \<inter> space M \<inter> A) = 0" by auto
   next
     fix x assume "x \<in> f`A \<union> {0}"
     hence "x \<noteq> 0 \<Longrightarrow> ?I -` {x} \<inter> space M = f -` {x} \<inter> space M \<inter> A"
       by (auto simp: indicator_def split: split_if_asm)
-    thus "x * \<mu> (?I -` {x} \<inter> space M) =
-      x * \<mu> (f -` {x} \<inter> space M \<inter> A)" by (cases "x = 0") simp_all
+    thus "x * (emeasure M) (?I -` {x} \<inter> space M) =
+      x * (emeasure M) (f -` {x} \<inter> space M \<inter> A)" by (cases "x = 0") simp_all
   qed
   show ?thesis unfolding *
     using assms(2) unfolding simple_function_def
     by (auto intro!: setsum_mono_zero_cong_right)
 qed
 
-lemma (in measure_space) simple_integral_indicator_only[simp]:
+lemma simple_integral_indicator_only[simp]:
   assumes "A \<in> sets M"
-  shows "integral\<^isup>S M (indicator A) = \<mu> A"
+  shows "integral\<^isup>S M (indicator A) = emeasure M A"
 proof cases
   assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto
   thus ?thesis unfolding simple_integral_def using `space M = {}` by auto
 next
   assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::ereal}" by auto
   thus ?thesis
-    using simple_integral_indicator[OF assms simple_function_const[of 1]]
+    using simple_integral_indicator[OF assms simple_function_const[of _ 1]]
     using sets_into_space[OF assms]
-    by (auto intro!: arg_cong[where f="\<mu>"])
+    by (auto intro!: arg_cong[where f="(emeasure M)"])
 qed
 
-lemma (in measure_space) simple_integral_null_set:
-  assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets"
+lemma simple_integral_null_set:
+  assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets M"
   shows "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = 0"
 proof -
-  have "AE x. indicator N x = (0 :: ereal)"
-    using `N \<in> null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N])
+  have "AE x in M. indicator N x = (0 :: ereal)"
+    using `N \<in> null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N])
   then have "(\<integral>\<^isup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^isup>Sx. 0 \<partial>M)"
     using assms apply (intro simple_integral_cong_AE) by auto
   then show ?thesis by simp
 qed
 
-lemma (in measure_space) simple_integral_cong_AE_mult_indicator:
-  assumes sf: "simple_function M f" and eq: "AE x. x \<in> S" and "S \<in> sets M"
+lemma simple_integral_cong_AE_mult_indicator:
+  assumes sf: "simple_function M f" and eq: "AE x in M. x \<in> S" and "S \<in> sets M"
   shows "integral\<^isup>S M f = (\<integral>\<^isup>Sx. f x * indicator S x \<partial>M)"
   using assms by (intro simple_integral_cong_AE) auto
 
-lemma (in measure_space) simple_integral_restricted:
-  assumes "A \<in> sets M"
-  assumes sf: "simple_function M (\<lambda>x. f x * indicator A x)"
-  shows "integral\<^isup>S (restricted_space A) f = (\<integral>\<^isup>Sx. f x * indicator A x \<partial>M)"
-    (is "_ = integral\<^isup>S M ?f")
+lemma simple_integral_distr:
+  assumes T: "T \<in> measurable M M'"
+    and f: "simple_function M' f"
+  shows "integral\<^isup>S (distr M M' T) f = (\<integral>\<^isup>S x. f (T x) \<partial>M)"
   unfolding simple_integral_def
-proof (simp, safe intro!: setsum_mono_zero_cong_left)
-  from sf show "finite (?f ` space M)"
-    unfolding simple_function_def by auto
+proof (intro setsum_mono_zero_cong_right ballI)
+  show "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space (distr M M' T)"
+    using T unfolding measurable_def by auto
+  show "finite (f ` space (distr M M' T))"
+    using f unfolding simple_function_def by auto
 next
-  fix x assume "x \<in> A"
-  then show "f x \<in> ?f ` space M"
-    using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x])
-next
-  fix x assume "x \<in> space M" "?f x \<notin> f`A"
-  then have "x \<notin> A" by (auto simp: image_iff)
-  then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp
+  fix i assume "i \<in> f ` space (distr M M' T) - (\<lambda>x. f (T x)) ` space M"
+  then have "T -` (f -` {i} \<inter> space (distr M M' T)) \<inter> space M = {}" by (auto simp: image_iff)
+  with f[THEN simple_functionD(2), of "{i}"]
+  show "i * emeasure (distr M M' T) (f -` {i} \<inter> space (distr M M' T)) = 0"
+    using T by (simp add: emeasure_distr)
 next
-  fix x assume "x \<in> A"
-  then have "f x \<noteq> 0 \<Longrightarrow>
-    f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M"
-    using `A \<in> sets M` sets_into_space
-    by (auto simp: indicator_def split: split_if_asm)
-  then show "f x * \<mu> (f -` {f x} \<inter> A) =
-    f x * \<mu> (?f -` {f x} \<inter> space M)"
-    unfolding ereal_mult_cancel_left by auto
+  fix i assume "i \<in> (\<lambda>x. f (T x)) ` space M"
+  then have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
+    using T unfolding measurable_def by auto
+  with f[THEN simple_functionD(2), of "{i}"] T
+  show "i * emeasure (distr M M' T) (f -` {i} \<inter> space (distr M M' T)) =
+      i * (emeasure M) ((\<lambda>x. f (T x)) -` {i} \<inter> space M)"
+    by (auto simp add: emeasure_distr)
 qed
 
-lemma (in measure_space) simple_integral_subalgebra:
-  assumes N: "measure_space N" and [simp]: "space N = space M" "measure N = measure M"
-  shows "integral\<^isup>S N = integral\<^isup>S M"
-  unfolding simple_integral_def [abs_def] by simp
-
-lemma (in measure_space) simple_integral_vimage:
-  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
-    and f: "simple_function M' f"
-  shows "integral\<^isup>S M' f = (\<integral>\<^isup>S x. f (T x) \<partial>M)"
-proof -
-  interpret T: measure_space M' by (rule measure_space_vimage[OF T])
-  show "integral\<^isup>S M' f = (\<integral>\<^isup>S x. f (T x) \<partial>M)"
-    unfolding simple_integral_def
-  proof (intro setsum_mono_zero_cong_right ballI)
-    show "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'"
-      using T unfolding measurable_def measure_preserving_def by auto
-    show "finite (f ` space M')"
-      using f unfolding simple_function_def by auto
-  next
-    fix i assume "i \<in> f ` space M' - (\<lambda>x. f (T x)) ` space M"
-    then have "T -` (f -` {i} \<inter> space M') \<inter> space M = {}" by (auto simp: image_iff)
-    with f[THEN T.simple_functionD(2), THEN measure_preservingD[OF T(2)], of "{i}"]
-    show "i * T.\<mu> (f -` {i} \<inter> space M') = 0" by simp
-  next
-    fix i assume "i \<in> (\<lambda>x. f (T x)) ` space M"
-    then have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
-      using T unfolding measurable_def measure_preserving_def by auto
-    with f[THEN T.simple_functionD(2), THEN measure_preservingD[OF T(2)], of "{i}"]
-    show "i * T.\<mu> (f -` {i} \<inter> space M') = i * \<mu> ((\<lambda>x. f (T x)) -` {i} \<inter> space M)"
-      by auto
-  qed
-qed
-
-lemma (in measure_space) simple_integral_cmult_indicator:
+lemma simple_integral_cmult_indicator:
   assumes A: "A \<in> sets M"
-  shows "(\<integral>\<^isup>Sx. c * indicator A x \<partial>M) = c * \<mu> A"
+  shows "(\<integral>\<^isup>Sx. c * indicator A x \<partial>M) = c * (emeasure M) A"
   using simple_integral_mult[OF simple_function_indicator[OF A]]
   unfolding simple_integral_indicator_only[OF A] by simp
 
-lemma (in measure_space) simple_integral_positive:
-  assumes f: "simple_function M f" and ae: "AE x. 0 \<le> f x"
+lemma simple_integral_positive:
+  assumes f: "simple_function M f" and ae: "AE x in M. 0 \<le> f x"
   shows "0 \<le> integral\<^isup>S M f"
 proof -
   have "integral\<^isup>S M (\<lambda>x. 0) \<le> integral\<^isup>S M f"
@@ -868,29 +772,23 @@
 
 section "Continuous positive integration"
 
-definition positive_integral_def:
+definition positive_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^isup>P") where
   "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^isup>S M g)"
 
 syntax
-  "_positive_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> ereal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
+  "_positive_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^isup>+ _. _ \<partial>_" [60,61] 110)
 
 translations
-  "\<integral>\<^isup>+ x. f \<partial>M" == "CONST integral\<^isup>P M (%x. f)"
+  "\<integral>\<^isup>+ x. f \<partial>M" == "CONST positive_integral M (%x. f)"
 
-lemma (in measure_space) positive_integral_cong_measure:
-  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
-  shows "integral\<^isup>P N f = integral\<^isup>P M f"
-  unfolding positive_integral_def
-  unfolding simple_function_cong_algebra[OF assms(2,3), symmetric]
-  using AE_cong_measure[OF assms]
-  using simple_integral_cong_measure[OF assms]
-  by (auto intro!: SUP_cong)
-
-lemma (in measure_space) positive_integral_positive:
+lemma positive_integral_positive:
   "0 \<le> integral\<^isup>P M f"
   by (auto intro!: SUP_upper2[of "\<lambda>x. 0"] simp: positive_integral_def le_fun_def)
 
-lemma (in measure_space) positive_integral_def_finite:
+lemma positive_integral_not_MInfty[simp]: "integral\<^isup>P M f \<noteq> -\<infinity>"
+  using positive_integral_positive[of M f] by auto
+
+lemma positive_integral_def_finite:
   "integral\<^isup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f \<and> range g \<subseteq> {0 ..< \<infinity>}}. integral\<^isup>S M g)"
     (is "_ = SUPR ?A ?f")
   unfolding positive_integral_def
@@ -898,7 +796,7 @@
   fix g assume g: "simple_function M g" "g \<le> max 0 \<circ> f"
   let ?G = "{x \<in> space M. \<not> g x \<noteq> \<infinity>}"
   note gM = g(1)[THEN borel_measurable_simple_function]
-  have \<mu>G_pos: "0 \<le> \<mu> ?G" using gM by auto
+  have \<mu>G_pos: "0 \<le> (emeasure M) ?G" using gM by auto
   let ?g = "\<lambda>y x. if g x = \<infinity> then y else max 0 (g x)"
   from g gM have g_in_A: "\<And>y. 0 \<le> y \<Longrightarrow> y \<noteq> \<infinity> \<Longrightarrow> ?g y \<in> ?A"
     apply (safe intro!: simple_function_max simple_function_If)
@@ -907,21 +805,22 @@
   show "integral\<^isup>S M g \<le> SUPR ?A ?f"
   proof cases
     have g0: "?g 0 \<in> ?A" by (intro g_in_A) auto
-    assume "\<mu> ?G = 0"
-    with gM have "AE x. x \<notin> ?G" by (simp add: AE_iff_null_set)
+    assume "(emeasure M) ?G = 0"
+    with gM have "AE x in M. x \<notin> ?G"
+      by (auto simp add: AE_iff_null intro!: null_setsI)
     with gM g show ?thesis
       by (intro SUP_upper2[OF g0] simple_integral_mono_AE)
          (auto simp: max_def intro!: simple_function_If)
   next
-    assume \<mu>G: "\<mu> ?G \<noteq> 0"
+    assume \<mu>G: "(emeasure M) ?G \<noteq> 0"
     have "SUPR ?A (integral\<^isup>S M) = \<infinity>"
     proof (intro SUP_PInfty)
       fix n :: nat
-      let ?y = "ereal (real n) / (if \<mu> ?G = \<infinity> then 1 else \<mu> ?G)"
+      let ?y = "ereal (real n) / (if (emeasure M) ?G = \<infinity> then 1 else (emeasure M) ?G)"
       have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>G \<mu>G_pos by (auto simp: ereal_divide_eq)
       then have "?g ?y \<in> ?A" by (rule g_in_A)
-      have "real n \<le> ?y * \<mu> ?G"
-        using \<mu>G \<mu>G_pos by (cases "\<mu> ?G") (auto simp: field_simps)
+      have "real n \<le> ?y * (emeasure M) ?G"
+        using \<mu>G \<mu>G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps)
       also have "\<dots> = (\<integral>\<^isup>Sx. ?y * indicator ?G x \<partial>M)"
         using `0 \<le> ?y` `?g ?y \<in> ?A` gM
         by (subst simple_integral_cmult_indicator) auto
@@ -934,15 +833,15 @@
   qed
 qed (auto intro: SUP_upper)
 
-lemma (in measure_space) positive_integral_mono_AE:
-  assumes ae: "AE x. u x \<le> v x" shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
+lemma positive_integral_mono_AE:
+  assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^isup>P M u \<le> integral\<^isup>P M v"
   unfolding positive_integral_def
 proof (safe intro!: SUP_mono)
   fix n assume n: "simple_function M n" "n \<le> max 0 \<circ> u"
   from ae[THEN AE_E] guess N . note N = this
-  then have ae_N: "AE x. x \<notin> N" by (auto intro: AE_not_in)
+  then have ae_N: "AE x in M. x \<notin> N" by (auto intro: AE_not_in)
   let ?n = "\<lambda>x. n x * indicator (space M - N) x"
-  have "AE x. n x \<le> ?n x" "simple_function M ?n"
+  have "AE x in M. n x \<le> ?n x" "simple_function M ?n"
     using n N ae_N by auto
   moreover
   { fix x have "?n x \<le> max 0 (v x)"
@@ -959,19 +858,19 @@
     by force
 qed
 
-lemma (in measure_space) positive_integral_mono:
+lemma positive_integral_mono:
   "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^isup>P M u \<le> integral\<^isup>P M v"
   by (auto intro: positive_integral_mono_AE)
 
-lemma (in measure_space) positive_integral_cong_AE:
-  "AE x. u x = v x \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
+lemma positive_integral_cong_AE:
+  "AE x in M. u x = v x \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
   by (auto simp: eq_iff intro!: positive_integral_mono_AE)
 
-lemma (in measure_space) positive_integral_cong:
+lemma positive_integral_cong:
   "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^isup>P M u = integral\<^isup>P M v"
   by (auto intro: positive_integral_cong_AE)
 
-lemma (in measure_space) positive_integral_eq_simple_integral:
+lemma positive_integral_eq_simple_integral:
   assumes f: "simple_function M f" "\<And>x. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
 proof -
   let ?f = "\<lambda>x. f x * indicator (space M) x"
@@ -987,10 +886,10 @@
     by (simp cong: positive_integral_cong simple_integral_cong)
 qed
 
-lemma (in measure_space) positive_integral_eq_simple_integral_AE:
-  assumes f: "simple_function M f" "AE x. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
+lemma positive_integral_eq_simple_integral_AE:
+  assumes f: "simple_function M f" "AE x in M. 0 \<le> f x" shows "integral\<^isup>P M f = integral\<^isup>S M f"
 proof -
-  have "AE x. f x = max 0 (f x)" using f by (auto split: split_max)
+  have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max)
   with f have "integral\<^isup>P M f = integral\<^isup>S M (\<lambda>x. max 0 (f x))"
     by (simp cong: positive_integral_cong_AE simple_integral_cong_AE
              add: positive_integral_eq_simple_integral)
@@ -998,7 +897,7 @@
     by (auto intro!: simple_integral_cong_AE split: split_max)
 qed
 
-lemma (in measure_space) positive_integral_SUP_approx:
+lemma positive_integral_SUP_approx:
   assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
   and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
   shows "integral\<^isup>S M u \<le> (SUP i. integral\<^isup>P M (f i))" (is "_ \<le> ?S")
@@ -1028,7 +927,7 @@
   note B_u = Int[OF u(1)[THEN simple_functionD(2)] B]
 
   let ?B' = "\<lambda>i n. (u -` {i} \<inter> space M) \<inter> ?B n"
-  have measure_conv: "\<And>i. \<mu> (u -` {i} \<inter> space M) = (SUP n. \<mu> (?B' i n))"
+  have measure_conv: "\<And>i. (emeasure M) (u -` {i} \<inter> space M) = (SUP n. (emeasure M) (?B' i n))"
   proof -
     fix i
     have 1: "range (?B' i) \<subseteq> sets M" using B_u by auto
@@ -1051,17 +950,17 @@
         thus ?thesis using `x \<in> space M` by auto
       qed
     qed
-    then show "?thesis i" using continuity_from_below[OF 1 2] by simp
+    then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp
   qed
 
   have "integral\<^isup>S M u = (SUP i. integral\<^isup>S M (?uB i))"
     unfolding simple_integral_indicator[OF B `simple_function M u`]
   proof (subst SUPR_ereal_setsum, safe)
     fix x n assume "x \<in> space M"
-    with u_range show "incseq (\<lambda>i. u x * \<mu> (?B' (u x) i))" "\<And>i. 0 \<le> u x * \<mu> (?B' (u x) i)"
-      using B_mono B_u by (auto intro!: measure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff)
+    with u_range show "incseq (\<lambda>i. u x * (emeasure M) (?B' (u x) i))" "\<And>i. 0 \<le> u x * (emeasure M) (?B' (u x) i)"
+      using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff)
   next
-    show "integral\<^isup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (?B' i n))"
+    show "integral\<^isup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * (emeasure M) (?B' i n))"
       using measure_conv u_range B_u unfolding simple_integral_def
       by (auto intro!: setsum_cong SUPR_ereal_cmult[symmetric])
   qed
@@ -1089,7 +988,7 @@
   ultimately show "a * integral\<^isup>S M u \<le> ?S" by simp
 qed
 
-lemma (in measure_space) incseq_positive_integral:
+lemma incseq_positive_integral:
   assumes "incseq f" shows "incseq (\<lambda>i. integral\<^isup>P M (f i))"
 proof -
   have "\<And>i x. f i x \<le> f (Suc i) x"
@@ -1099,7 +998,7 @@
 qed
 
 text {* Beppo-Levi monotone convergence theorem *}
-lemma (in measure_space) positive_integral_monotone_convergence_SUP:
+lemma positive_integral_monotone_convergence_SUP:
   assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
   shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
 proof (rule antisym)
@@ -1107,7 +1006,7 @@
     by (auto intro!: SUP_least SUP_upper positive_integral_mono)
 next
   show "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP j. integral\<^isup>P M (f j))"
-    unfolding positive_integral_def_finite[of "\<lambda>x. SUP i. f i x"]
+    unfolding positive_integral_def_finite[of _ "\<lambda>x. SUP i. f i x"]
   proof (safe intro!: SUP_least)
     fix g assume g: "simple_function M g"
       and "g \<le> max 0 \<circ> (\<lambda>x. SUP i. f i x)" "range g \<subseteq> {0..<\<infinity>}"
@@ -1119,15 +1018,15 @@
   qed
 qed
 
-lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE:
-  assumes f: "\<And>i. AE x. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x" "\<And>i. f i \<in> borel_measurable M"
+lemma positive_integral_monotone_convergence_SUP_AE:
+  assumes f: "\<And>i. AE x in M. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x" "\<And>i. f i \<in> borel_measurable M"
   shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
 proof -
-  from f have "AE x. \<forall>i. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x"
+  from f have "AE x in M. \<forall>i. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x"
     by (simp add: AE_all_countable)
   from this[THEN AE_E] guess N . note N = this
   let ?f = "\<lambda>i x. if x \<in> space M - N then f i x else 0"
-  have f_eq: "AE x. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ N])
+  have f_eq: "AE x in M. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N])
   then have "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^isup>+ x. (SUP i. ?f i x) \<partial>M)"
     by (auto intro!: positive_integral_cong_AE)
   also have "\<dots> = (SUP i. (\<integral>\<^isup>+ x. ?f i x \<partial>M))"
@@ -1143,14 +1042,14 @@
   finally show ?thesis .
 qed
 
-lemma (in measure_space) positive_integral_monotone_convergence_SUP_AE_incseq:
-  assumes f: "incseq f" "\<And>i. AE x. 0 \<le> f i x" and borel: "\<And>i. f i \<in> borel_measurable M"
+lemma positive_integral_monotone_convergence_SUP_AE_incseq:
+  assumes f: "incseq f" "\<And>i. AE x in M. 0 \<le> f i x" and borel: "\<And>i. f i \<in> borel_measurable M"
   shows "(\<integral>\<^isup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^isup>P M (f i))"
   using f[unfolded incseq_Suc_iff le_fun_def]
   by (intro positive_integral_monotone_convergence_SUP_AE[OF _ borel])
      auto
 
-lemma (in measure_space) positive_integral_monotone_convergence_simple:
+lemma positive_integral_monotone_convergence_simple:
   assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
   shows "(SUP i. integral\<^isup>S M (f i)) = (\<integral>\<^isup>+x. (SUP i. f i x) \<partial>M)"
   using assms unfolding positive_integral_monotone_convergence_SUP[OF f(1)
@@ -1161,7 +1060,7 @@
   "(\<integral>\<^isup>+x. max 0 (f x) \<partial>M) = integral\<^isup>P M f"
   by (simp add: le_fun_def positive_integral_def)
 
-lemma (in measure_space) positive_integral_cong_pos:
+lemma positive_integral_cong_pos:
   assumes "\<And>x. x \<in> space M \<Longrightarrow> f x \<le> 0 \<and> g x \<le> 0 \<or> f x = g x"
   shows "integral\<^isup>P M f = integral\<^isup>P M g"
 proof -
@@ -1174,10 +1073,10 @@
   then show ?thesis by (simp add: positive_integral_max_0)
 qed
 
-lemma (in measure_space) SUP_simple_integral_sequences:
+lemma SUP_simple_integral_sequences:
   assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)"
   and g: "incseq g" "\<And>i x. 0 \<le> g i x" "\<And>i. simple_function M (g i)"
-  and eq: "AE x. (SUP i. f i x) = (SUP i. g i x)"
+  and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)"
   shows "(SUP i. integral\<^isup>S M (f i)) = (SUP i. integral\<^isup>S M (g i))"
     (is "SUPR _ ?F = SUPR _ ?G")
 proof -
@@ -1190,32 +1089,11 @@
   finally show ?thesis by simp
 qed
 
-lemma (in measure_space) positive_integral_const[simp]:
-  "0 \<le> c \<Longrightarrow> (\<integral>\<^isup>+ x. c \<partial>M) = c * \<mu> (space M)"
+lemma positive_integral_const[simp]:
+  "0 \<le> c \<Longrightarrow> (\<integral>\<^isup>+ x. c \<partial>M) = c * (emeasure M) (space M)"
   by (subst positive_integral_eq_simple_integral) auto
 
-lemma (in measure_space) positive_integral_vimage:
-  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
-  and f: "f \<in> borel_measurable M'"
-  shows "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
-proof -
-  interpret T: measure_space M' by (rule measure_space_vimage[OF T])
-  from T.borel_measurable_implies_simple_function_sequence'[OF f]
-  guess f' . note f' = this
-  let ?f = "\<lambda>i x. f' i (T x)"
-  have inc: "incseq ?f" using f' by (force simp: le_fun_def incseq_def)
-  have sup: "\<And>x. (SUP i. ?f i x) = max 0 (f (T x))"
-    using f'(4) .
-  have sf: "\<And>i. simple_function M (\<lambda>x. f' i (T x))"
-    using simple_function_vimage[OF T(1) measure_preservingD2[OF T(2)] f'(1)] .
-  show "integral\<^isup>P M' f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
-    using
-      T.positive_integral_monotone_convergence_simple[OF f'(2,5,1)]
-      positive_integral_monotone_convergence_simple[OF inc f'(5) sf]
-    by (simp add: positive_integral_max_0 simple_integral_vimage[OF T f'(1)] f')
-qed
-
-lemma (in measure_space) positive_integral_linear:
+lemma positive_integral_linear:
   assumes f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and "0 \<le> a"
   and g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x"
   shows "(\<integral>\<^isup>+ x. a * f x + g x \<partial>M) = a * integral\<^isup>P M f + integral\<^isup>P M g"
@@ -1254,7 +1132,7 @@
         by (subst SUPR_ereal_cmult[symmetric, OF u(6) `0 \<le> a`])
            (auto intro!: SUPR_ereal_add
                  simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono ereal_add_nonneg_nonneg) }
-    then show "AE x. (SUP i. l i x) = (SUP i. ?L' i x)"
+    then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)"
       unfolding l(5) using `0 \<le> a` u(5) v(5) l(5) f(2) g(2)
       by (intro AE_I2) (auto split: split_max simp add: ereal_add_nonneg_nonneg)
   qed
@@ -1268,8 +1146,8 @@
   then show ?thesis by (simp add: positive_integral_max_0)
 qed
 
-lemma (in measure_space) positive_integral_cmult:
-  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x" "0 \<le> c"
+lemma positive_integral_cmult:
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> c"
   shows "(\<integral>\<^isup>+ x. c * f x \<partial>M) = c * integral\<^isup>P M f"
 proof -
   have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c`
@@ -1277,68 +1155,68 @@
   have "(\<integral>\<^isup>+ x. c * f x \<partial>M) = (\<integral>\<^isup>+ x. c * max 0 (f x) \<partial>M)"
     by (simp add: positive_integral_max_0)
   then show ?thesis
-    using positive_integral_linear[OF _ _ `0 \<le> c`, of "\<lambda>x. max 0 (f x)" "\<lambda>x. 0"] f
+    using positive_integral_linear[OF _ _ `0 \<le> c`, of "\<lambda>x. max 0 (f x)" _ "\<lambda>x. 0"] f
     by (auto simp: positive_integral_max_0)
 qed
 
-lemma (in measure_space) positive_integral_multc:
-  assumes "f \<in> borel_measurable M" "AE x. 0 \<le> f x" "0 \<le> c"
+lemma positive_integral_multc:
+  assumes "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" "0 \<le> c"
   shows "(\<integral>\<^isup>+ x. f x * c \<partial>M) = integral\<^isup>P M f * c"
   unfolding mult_commute[of _ c] positive_integral_cmult[OF assms] by simp
 
-lemma (in measure_space) positive_integral_indicator[simp]:
-  "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. indicator A x\<partial>M) = \<mu> A"
+lemma positive_integral_indicator[simp]:
+  "A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. indicator A x\<partial>M) = (emeasure M) A"
   by (subst positive_integral_eq_simple_integral)
      (auto simp: simple_function_indicator simple_integral_indicator)
 
-lemma (in measure_space) positive_integral_cmult_indicator:
-  "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. c * indicator A x \<partial>M) = c * \<mu> A"
+lemma positive_integral_cmult_indicator:
+  "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+ x. c * indicator A x \<partial>M) = c * (emeasure M) A"
   by (subst positive_integral_eq_simple_integral)
      (auto simp: simple_function_indicator simple_integral_indicator)
 
-lemma (in measure_space) positive_integral_add:
-  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
-  and g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
+lemma positive_integral_add:
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
   shows "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = integral\<^isup>P M f + integral\<^isup>P M g"
 proof -
-  have ae: "AE x. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)"
+  have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)"
     using assms by (auto split: split_max simp: ereal_add_nonneg_nonneg)
   have "(\<integral>\<^isup>+ x. f x + g x \<partial>M) = (\<integral>\<^isup>+ x. max 0 (f x + g x) \<partial>M)"
     by (simp add: positive_integral_max_0)
   also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)"
     unfolding ae[THEN positive_integral_cong_AE] ..
   also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) \<partial>M) + (\<integral>\<^isup>+ x. max 0 (g x) \<partial>M)"
-    using positive_integral_linear[of "\<lambda>x. max 0 (f x)" 1 "\<lambda>x. max 0 (g x)"] f g
+    using positive_integral_linear[of "\<lambda>x. max 0 (f x)" _ 1 "\<lambda>x. max 0 (g x)"] f g
     by auto
   finally show ?thesis
     by (simp add: positive_integral_max_0)
 qed
 
-lemma (in measure_space) positive_integral_setsum:
-  assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M" "\<And>i. i\<in>P \<Longrightarrow> AE x. 0 \<le> f i x"
+lemma positive_integral_setsum:
+  assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M" "\<And>i. i\<in>P \<Longrightarrow> AE x in M. 0 \<le> f i x"
   shows "(\<integral>\<^isup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^isup>P M (f i))"
 proof cases
   assume f: "finite P"
-  from assms have "AE x. \<forall>i\<in>P. 0 \<le> f i x" unfolding AE_finite_all[OF f] by auto
+  from assms have "AE x in M. \<forall>i\<in>P. 0 \<le> f i x" unfolding AE_finite_all[OF f] by auto
   from f this assms(1) show ?thesis
   proof induct
     case (insert i P)
-    then have "f i \<in> borel_measurable M" "AE x. 0 \<le> f i x"
-      "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" "AE x. 0 \<le> (\<Sum>i\<in>P. f i x)"
+    then have "f i \<in> borel_measurable M" "AE x in M. 0 \<le> f i x"
+      "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" "AE x in M. 0 \<le> (\<Sum>i\<in>P. f i x)"
       by (auto intro!: borel_measurable_ereal_setsum setsum_nonneg)
     from positive_integral_add[OF this]
     show ?case using insert by auto
   qed simp
 qed simp
 
-lemma (in measure_space) positive_integral_Markov_inequality:
-  assumes u: "u \<in> borel_measurable M" "AE x. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c" "c \<noteq> \<infinity>"
-  shows "\<mu> ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
-    (is "\<mu> ?A \<le> _ * ?PI")
+lemma positive_integral_Markov_inequality:
+  assumes u: "u \<in> borel_measurable M" "AE x in M. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c" "c \<noteq> \<infinity>"
+  shows "(emeasure M) ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)"
+    (is "(emeasure M) ?A \<le> _ * ?PI")
 proof -
   have "?A \<in> sets M"
     using `A \<in> sets M` u by auto
-  hence "\<mu> ?A = (\<integral>\<^isup>+ x. indicator ?A x \<partial>M)"
+  hence "(emeasure M) ?A = (\<integral>\<^isup>+ x. indicator ?A x \<partial>M)"
     using positive_integral_indicator by simp
   also have "\<dots> \<le> (\<integral>\<^isup>+ x. c * (u x * indicator A x) \<partial>M)" using u c
     by (auto intro!: positive_integral_mono_AE
@@ -1349,17 +1227,17 @@
   finally show ?thesis .
 qed
 
-lemma (in measure_space) positive_integral_noteq_infinite:
-  assumes g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
+lemma positive_integral_noteq_infinite:
+  assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
   and "integral\<^isup>P M g \<noteq> \<infinity>"
-  shows "AE x. g x \<noteq> \<infinity>"
+  shows "AE x in M. g x \<noteq> \<infinity>"
 proof (rule ccontr)
-  assume c: "\<not> (AE x. g x \<noteq> \<infinity>)"
-  have "\<mu> {x\<in>space M. g x = \<infinity>} \<noteq> 0"
-    using c g by (simp add: AE_iff_null_set)
-  moreover have "0 \<le> \<mu> {x\<in>space M. g x = \<infinity>}" using g by (auto intro: measurable_sets)
-  ultimately have "0 < \<mu> {x\<in>space M. g x = \<infinity>}" by auto
-  then have "\<infinity> = \<infinity> * \<mu> {x\<in>space M. g x = \<infinity>}" by auto
+  assume c: "\<not> (AE x in M. g x \<noteq> \<infinity>)"
+  have "(emeasure M) {x\<in>space M. g x = \<infinity>} \<noteq> 0"
+    using c g by (auto simp add: AE_iff_null)
+  moreover have "0 \<le> (emeasure M) {x\<in>space M. g x = \<infinity>}" using g by (auto intro: measurable_sets)
+  ultimately have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}" by auto
+  then have "\<infinity> = \<infinity> * (emeasure M) {x\<in>space M. g x = \<infinity>}" by auto
   also have "\<dots> \<le> (\<integral>\<^isup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
     using g by (subst positive_integral_cmult_indicator) auto
   also have "\<dots> \<le> integral\<^isup>P M g"
@@ -1367,34 +1245,34 @@
   finally show False using `integral\<^isup>P M g \<noteq> \<infinity>` by auto
 qed
 
-lemma (in measure_space) positive_integral_diff:
+lemma positive_integral_diff:
   assumes f: "f \<in> borel_measurable M"
-  and g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
+  and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
   and fin: "integral\<^isup>P M g \<noteq> \<infinity>"
-  and mono: "AE x. g x \<le> f x"
+  and mono: "AE x in M. g x \<le> f x"
   shows "(\<integral>\<^isup>+ x. f x - g x \<partial>M) = integral\<^isup>P M f - integral\<^isup>P M g"
 proof -
-  have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x. 0 \<le> f x - g x"
+  have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x in M. 0 \<le> f x - g x"
     using assms by (auto intro: ereal_diff_positive)
-  have pos_f: "AE x. 0 \<le> f x" using mono g by auto
+  have pos_f: "AE x in M. 0 \<le> f x" using mono g by auto
   { fix a b :: ereal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
       by (cases rule: ereal2_cases[of a b]) auto }
   note * = this
-  then have "AE x. f x = f x - g x + g x"
+  then have "AE x in M. f x = f x - g x + g x"
     using mono positive_integral_noteq_infinite[OF g fin] assms by auto
   then have **: "integral\<^isup>P M f = (\<integral>\<^isup>+x. f x - g x \<partial>M) + integral\<^isup>P M g"
     unfolding positive_integral_add[OF diff g, symmetric]
     by (rule positive_integral_cong_AE)
   show ?thesis unfolding **
-    using fin positive_integral_positive[of g]
+    using fin positive_integral_positive[of M g]
     by (cases rule: ereal2_cases[of "\<integral>\<^isup>+ x. f x - g x \<partial>M" "integral\<^isup>P M g"]) auto
 qed
 
-lemma (in measure_space) positive_integral_suminf:
-  assumes f: "\<And>i. f i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> f i x"
+lemma positive_integral_suminf:
+  assumes f: "\<And>i. f i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> f i x"
   shows "(\<integral>\<^isup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^isup>P M (f i))"
 proof -
-  have all_pos: "AE x. \<forall>i. 0 \<le> f i x"
+  have all_pos: "AE x in M. \<forall>i. 0 \<le> f i x"
     using assms by (auto simp: AE_all_countable)
   have "(\<Sum>i. integral\<^isup>P M (f i)) = (SUP n. \<Sum>i<n. integral\<^isup>P M (f i))"
     using positive_integral_positive by (rule suminf_ereal_eq_SUPR)
@@ -1409,12 +1287,12 @@
 qed
 
 text {* Fatou's lemma: convergence theorem on limes inferior *}
-lemma (in measure_space) positive_integral_lim_INF:
+lemma positive_integral_lim_INF:
   fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
-  assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x. 0 \<le> u i x"
+  assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> u i x"
   shows "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^isup>P M (u n))"
 proof -
-  have pos: "AE x. \<forall>i. 0 \<le> u i x" using u by (auto simp: AE_all_countable)
+  have pos: "AE x in M. \<forall>i. 0 \<le> u i x" using u by (auto simp: AE_all_countable)
   have "(\<integral>\<^isup>+ x. liminf (\<lambda>n. u n x) \<partial>M) =
     (SUP n. \<integral>\<^isup>+ x. (INF i:{n..}. u i x) \<partial>M)"
     unfolding liminf_SUPR_INFI using pos u
@@ -1426,120 +1304,31 @@
   finally show ?thesis .
 qed
 
-lemma (in measure_space) measure_space_density:
-  assumes u: "u \<in> borel_measurable M" "AE x. 0 \<le> u x"
-    and M'[simp]: "M' = (M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+ x. u x * indicator A x \<partial>M)\<rparr>)"
-  shows "measure_space M'"
-proof -
-  interpret M': sigma_algebra M' by (intro sigma_algebra_cong) auto
-  show ?thesis
-  proof
-    have pos: "\<And>A. AE x. 0 \<le> u x * indicator A x"
-      using u by (auto simp: ereal_zero_le_0_iff)
-    then show "positive M' (measure M')" unfolding M'
-      using u(1) by (auto simp: positive_def intro!: positive_integral_positive)
-    show "countably_additive M' (measure M')"
-    proof (intro countably_additiveI)
-      fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M'"
-      then have *: "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M"
-        using u by (auto intro: borel_measurable_indicator)
-      assume disj: "disjoint_family A"
-      have "(\<Sum>n. measure M' (A n)) = (\<integral>\<^isup>+ x. (\<Sum>n. u x * indicator (A n) x) \<partial>M)"
-        unfolding M' using u(1) *
-        by (simp add: positive_integral_suminf[OF _ pos, symmetric])
-      also have "\<dots> = (\<integral>\<^isup>+ x. u x * (\<Sum>n. indicator (A n) x) \<partial>M)" using u
-        by (intro positive_integral_cong_AE)
-           (elim AE_mp, auto intro!: AE_I2 suminf_cmult_ereal)
-      also have "\<dots> = (\<integral>\<^isup>+ x. u x * indicator (\<Union>n. A n) x \<partial>M)"
-        unfolding suminf_indicator[OF disj] ..
-      finally show "(\<Sum>n. measure M' (A n)) = measure M' (\<Union>x. A x)"
-        unfolding M' by simp
-    qed
-  qed
-qed
-
-lemma (in measure_space) positive_integral_null_set:
-  assumes "N \<in> null_sets" shows "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = 0"
+lemma positive_integral_null_set:
+  assumes "N \<in> null_sets M" shows "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = 0"
 proof -
   have "(\<integral>\<^isup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^isup>+ x. 0 \<partial>M)"
   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"
+    show "(emeasure M) N = 0" "N \<in> sets M"
       using assms by auto
   qed
   then show ?thesis by simp
 qed
 
-lemma (in measure_space) positive_integral_translated_density:
-  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
-  assumes g: "g \<in> borel_measurable M" "AE x. 0 \<le> g x"
-    and M': "M' = (M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)\<rparr>)"
-  shows "integral\<^isup>P M' g = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
-proof -
-  from measure_space_density[OF f M']
-  interpret T: measure_space M' .
-  have borel[simp]:
-    "borel_measurable M' = borel_measurable M"
-    "simple_function M' = simple_function M"
-    unfolding measurable_def simple_function_def [abs_def] by (auto simp: M')
-  from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess G . note G = this
-  note G' = borel_measurable_simple_function[OF this(1)] simple_functionD[OF G(1)]
-  note G'(2)[simp]
-  { fix P have "AE x. P x \<Longrightarrow> AE x in M'. P x"
-      using positive_integral_null_set[of _ f]
-      unfolding T.almost_everywhere_def almost_everywhere_def
-      by (auto simp: M') }
-  note ac = this
-  from G(4) g(2) have G_M': "AE x in M'. (SUP i. G i x) = g x"
-    by (auto intro!: ac split: split_max)
-  { fix i
-    let ?I = "\<lambda>y x. indicator (G i -` {y} \<inter> space M) x"
-    { fix x assume *: "x \<in> space M" "0 \<le> f x" "0 \<le> g x"
-      then have [simp]: "G i ` space M \<inter> {y. G i x = y \<and> x \<in> space M} = {G i x}" by auto
-      from * G' G have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * (\<Sum>y\<in>G i`space M. (y * ?I y x))"
-        by (subst setsum_ereal_right_distrib) (auto simp: ac_simps)
-      also have "\<dots> = f x * G i x"
-        by (simp add: indicator_def if_distrib setsum_cases)
-      finally have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * G i x" . }
-    note to_singleton = this
-    have "integral\<^isup>P M' (G i) = integral\<^isup>S M' (G i)"
-      using G T.positive_integral_eq_simple_integral by simp
-    also have "\<dots> = (\<Sum>y\<in>G i`space M. y * (\<integral>\<^isup>+x. f x * ?I y x \<partial>M))"
-      unfolding simple_integral_def M' by simp
-    also have "\<dots> = (\<Sum>y\<in>G i`space M. (\<integral>\<^isup>+x. y * (f x * ?I y x) \<partial>M))"
-      using f G' G by (auto intro!: setsum_cong positive_integral_cmult[symmetric])
-    also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) \<partial>M)"
-      using f G' G by (auto intro!: positive_integral_setsum[symmetric])
-    finally have "integral\<^isup>P M' (G i) = (\<integral>\<^isup>+x. f x * G i x \<partial>M)"
-      using f g G' to_singleton by (auto intro!: positive_integral_cong_AE) }
-  note [simp] = this
-  have "integral\<^isup>P M' g = (SUP i. integral\<^isup>P M' (G i))" using G'(1) G_M'(1) G
-    using T.positive_integral_monotone_convergence_SUP[symmetric, OF `incseq G`]
-    by (simp cong: T.positive_integral_cong_AE)
-  also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f x * G i x \<partial>M))" by simp
-  also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. f x * G i x) \<partial>M)"
-    using f G' G(2)[THEN incseq_SucD] G
-    by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
-       (auto simp: ereal_mult_left_mono le_fun_def ereal_zero_le_0_iff)
-  also have "\<dots> = (\<integral>\<^isup>+x. f x * g x \<partial>M)" using f G' G g
-    by (intro positive_integral_cong_AE)
-       (auto simp add: SUPR_ereal_cmult split: split_max)
-  finally show "integral\<^isup>P M' g = (\<integral>\<^isup>+x. f x * g x \<partial>M)" .
-qed
-
-lemma (in measure_space) positive_integral_0_iff:
-  assumes u: "u \<in> borel_measurable M" and pos: "AE x. 0 \<le> u x"
-  shows "integral\<^isup>P M u = 0 \<longleftrightarrow> \<mu> {x\<in>space M. u x \<noteq> 0} = 0"
-    (is "_ \<longleftrightarrow> \<mu> ?A = 0")
+lemma positive_integral_0_iff:
+  assumes u: "u \<in> borel_measurable M" and pos: "AE x in M. 0 \<le> u x"
+  shows "integral\<^isup>P M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
+    (is "_ \<longleftrightarrow> (emeasure M) ?A = 0")
 proof -
   have u_eq: "(\<integral>\<^isup>+ x. u x * indicator ?A x \<partial>M) = integral\<^isup>P M u"
     by (auto intro!: positive_integral_cong simp: indicator_def)
   show ?thesis
   proof
-    assume "\<mu> ?A = 0"
-    with positive_integral_null_set[of ?A u] u
-    show "integral\<^isup>P M u = 0" by (simp add: u_eq)
+    assume "(emeasure M) ?A = 0"
+    with positive_integral_null_set[of ?A M u] u
+    show "integral\<^isup>P M u = 0" by (simp add: u_eq null_sets_def)
   next
     { fix r :: ereal and n :: nat assume gt_1: "1 \<le> real n * r"
       then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def)
@@ -1547,17 +1336,17 @@
     note gt_1 = this
     assume *: "integral\<^isup>P M u = 0"
     let ?M = "\<lambda>n. {x \<in> space M. 1 \<le> real (n::nat) * u x}"
-    have "0 = (SUP n. \<mu> (?M n \<inter> ?A))"
+    have "0 = (SUP n. (emeasure M) (?M n \<inter> ?A))"
     proof -
       { fix n :: nat
         from positive_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"]
-        have "\<mu> (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp
-        moreover have "0 \<le> \<mu> (?M n \<inter> ?A)" using u by auto
-        ultimately have "\<mu> (?M n \<inter> ?A) = 0" by auto }
+        have "(emeasure M) (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp
+        moreover have "0 \<le> (emeasure M) (?M n \<inter> ?A)" using u by auto
+        ultimately have "(emeasure M) (?M n \<inter> ?A) = 0" by auto }
       thus ?thesis by simp
     qed
-    also have "\<dots> = \<mu> (\<Union>n. ?M n \<inter> ?A)"
-    proof (safe intro!: continuity_from_below)
+    also have "\<dots> = (emeasure M) (\<Union>n. ?M n \<inter> ?A)"
+    proof (safe intro!: SUP_emeasure_incseq)
       fix n show "?M n \<inter> ?A \<in> sets M"
         using u by (auto intro!: Int)
     next
@@ -1570,8 +1359,8 @@
         finally show "1 \<le> real (Suc n) * u x" by auto
       qed
     qed
-    also have "\<dots> = \<mu> {x\<in>space M. 0 < u x}"
-    proof (safe intro!: arg_cong[where f="\<mu>"] dest!: gt_1)
+    also have "\<dots> = (emeasure M) {x\<in>space M. 0 < u x}"
+    proof (safe intro!: arg_cong[where f="(emeasure M)"] dest!: gt_1)
       fix x assume "0 < u x" and [simp, intro]: "x \<in> space M"
       show "x \<in> (\<Union>n. ?M n \<inter> ?A)"
       proof (cases "u x")
@@ -1582,88 +1371,48 @@
         thus ?thesis using `0 < r` real by (auto simp: one_ereal_def)
       qed (insert `0 < u x`, auto)
     qed auto
-    finally have "\<mu> {x\<in>space M. 0 < u x} = 0" by simp
+    finally have "(emeasure M) {x\<in>space M. 0 < u x} = 0" by simp
     moreover
-    from pos have "AE x. \<not> (u x < 0)" by auto
-    then have "\<mu> {x\<in>space M. u x < 0} = 0"
-      using AE_iff_null_set u by auto
-    moreover have "\<mu> {x\<in>space M. u x \<noteq> 0} = \<mu> {x\<in>space M. u x < 0} + \<mu> {x\<in>space M. 0 < u x}"
-      using u by (subst measure_additive) (auto intro!: arg_cong[where f=\<mu>])
-    ultimately show "\<mu> ?A = 0" by simp
+    from pos have "AE x in M. \<not> (u x < 0)" by auto
+    then have "(emeasure M) {x\<in>space M. u x < 0} = 0"
+      using AE_iff_null[of M] u by auto
+    moreover have "(emeasure M) {x\<in>space M. u x \<noteq> 0} = (emeasure M) {x\<in>space M. u x < 0} + (emeasure M) {x\<in>space M. 0 < u x}"
+      using u by (subst plus_emeasure) (auto intro!: arg_cong[where f="emeasure M"])
+    ultimately show "(emeasure M) ?A = 0" by simp
   qed
 qed
 
-lemma (in measure_space) positive_integral_0_iff_AE:
+lemma positive_integral_0_iff_AE:
   assumes u: "u \<in> borel_measurable M"
-  shows "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. u x \<le> 0)"
+  shows "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x in M. u x \<le> 0)"
 proof -
   have sets: "{x\<in>space M. max 0 (u x) \<noteq> 0} \<in> sets M"
     using u by auto
   from positive_integral_0_iff[of "\<lambda>x. max 0 (u x)"]
-  have "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x. max 0 (u x) = 0)"
+  have "integral\<^isup>P M u = 0 \<longleftrightarrow> (AE x in M. max 0 (u x) = 0)"
     unfolding positive_integral_max_0
-    using AE_iff_null_set[OF sets] u by auto
-  also have "\<dots> \<longleftrightarrow> (AE x. u x \<le> 0)" by (auto split: split_max)
+    using AE_iff_null[OF sets] u by auto
+  also have "\<dots> \<longleftrightarrow> (AE x in M. u x \<le> 0)" by (auto split: split_max)
   finally show ?thesis .
 qed
 
-lemma (in measure_space) positive_integral_const_If:
-  "(\<integral>\<^isup>+x. a \<partial>M) = (if 0 \<le> a then a * \<mu> (space M) else 0)"
+lemma positive_integral_const_If:
+  "(\<integral>\<^isup>+x. a \<partial>M) = (if 0 \<le> a then a * (emeasure M) (space M) else 0)"
   by (auto intro!: positive_integral_0_iff_AE[THEN iffD2])
 
-lemma (in measure_space) positive_integral_restricted:
-  assumes A: "A \<in> sets M"
-  shows "integral\<^isup>P (restricted_space A) f = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
-    (is "integral\<^isup>P ?R f = integral\<^isup>P M ?f")
-proof -
-  interpret R: measure_space ?R
-    by (rule restricted_measure_space) fact
-  let ?I = "\<lambda>g x. g x * indicator A x :: ereal"
-  show ?thesis
-    unfolding positive_integral_def
-    unfolding simple_function_restricted[OF A]
-    unfolding AE_restricted[OF A]
-  proof (safe intro!: SUPR_eq)
-    fix g assume g: "simple_function M (?I g)" and le: "g \<le> max 0 \<circ> f"
-    show "\<exists>j\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> ?I f}.
-      integral\<^isup>S (restricted_space A) g \<le> integral\<^isup>S M j"
-    proof (safe intro!: bexI[of _ "?I g"])
-      show "integral\<^isup>S (restricted_space A) g \<le> integral\<^isup>S M (?I g)"
-        using g A by (simp add: simple_integral_restricted)
-      show "?I g \<le> max 0 \<circ> ?I f"
-        using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
-    qed fact
-  next
-    fix g assume g: "simple_function M g" and le: "g \<le> max 0 \<circ> ?I f"
-    show "\<exists>i\<in>{g. simple_function M (?I g) \<and> g \<le> max 0 \<circ> f}.
-      integral\<^isup>S M g \<le> integral\<^isup>S (restricted_space A) i"
-    proof (safe intro!: bexI[of _ "?I g"])
-      show "?I g \<le> max 0 \<circ> f"
-        using le by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
-      from le have "\<And>x. g x \<le> ?I (?I g) x"
-        by (auto simp: le_fun_def max_def indicator_def split: split_if_asm)
-      then show "integral\<^isup>S M g \<le> integral\<^isup>S (restricted_space A) (?I g)"
-        using A g by (auto intro!: simple_integral_mono simp: simple_integral_restricted)
-      show "simple_function M (?I (?I g))" using g A by auto
-    qed
-  qed
-qed
-
-lemma (in measure_space) positive_integral_subalgebra:
+lemma positive_integral_subalgebra:
   assumes f: "f \<in> borel_measurable N" "AE x in N. 0 \<le> f x"
-  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
-  and sa: "sigma_algebra N"
+  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
   shows "integral\<^isup>P N f = integral\<^isup>P M f"
 proof -
-  interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
-  from N.borel_measurable_implies_simple_function_sequence'[OF f(1)] guess fs . note fs = this
+  from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess fs . note fs = this
   note sf = simple_function_subalgebra[OF fs(1) N(1,2)]
-  from N.positive_integral_monotone_convergence_simple[OF fs(2,5,1), symmetric]
-  have "integral\<^isup>P N f = (SUP i. \<Sum>x\<in>fs i ` space M. x * N.\<mu> (fs i -` {x} \<inter> space M))"
+  from positive_integral_monotone_convergence_simple[OF fs(2,5,1), symmetric]
+  have "integral\<^isup>P N f = (SUP i. \<Sum>x\<in>fs i ` space M. x * emeasure N (fs i -` {x} \<inter> space M))"
     unfolding fs(4) positive_integral_max_0
     unfolding simple_integral_def `space N = space M` by simp
-  also have "\<dots> = (SUP i. \<Sum>x\<in>fs i ` space M. x * \<mu> (fs i -` {x} \<inter> space M))"
-    using N N.simple_functionD(2)[OF fs(1)] unfolding `space N = space M` by auto
+  also have "\<dots> = (SUP i. \<Sum>x\<in>fs i ` space M. x * (emeasure M) (fs i -` {x} \<inter> space M))"
+    using N simple_functionD(2)[OF fs(1)] unfolding `space N = space M` by auto
   also have "\<dots> = integral\<^isup>P M f"
     using positive_integral_monotone_convergence_simple[OF fs(2,5) sf, symmetric]
     unfolding fs(4) positive_integral_max_0
@@ -1673,7 +1422,7 @@
 
 section "Lebesgue Integral"
 
-definition integrable where
+definition integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> bool" where
   "integrable M f \<longleftrightarrow> f \<in> borel_measurable M \<and>
     (\<integral>\<^isup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity> \<and> (\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
 
@@ -1682,55 +1431,44 @@
   shows "f \<in> borel_measurable M" "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) \<noteq> \<infinity>" "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>M) \<noteq> \<infinity>"
   using assms unfolding integrable_def by auto
 
-definition lebesgue_integral_def:
+definition lebesgue_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> real) \<Rightarrow> real" ("integral\<^isup>L") where
   "integral\<^isup>L M f = real ((\<integral>\<^isup>+ x. ereal (f x) \<partial>M)) - real ((\<integral>\<^isup>+ x. ereal (- f x) \<partial>M))"
 
 syntax
-  "_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
+  "_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> 'a measure \<Rightarrow> real" ("\<integral> _. _ \<partial>_" [60,61] 110)
 
 translations
-  "\<integral> x. f \<partial>M" == "CONST integral\<^isup>L M (%x. f)"
+  "\<integral> x. f \<partial>M" == "CONST lebesgue_integral M (%x. f)"
 
-lemma (in measure_space) integrableE:
+lemma integrableE:
   assumes "integrable M f"
   obtains r q where
     "(\<integral>\<^isup>+x. ereal (f x)\<partial>M) = ereal r"
     "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M) = ereal q"
     "f \<in> borel_measurable M" "integral\<^isup>L M f = r - q"
   using assms unfolding integrable_def lebesgue_integral_def
-  using positive_integral_positive[of "\<lambda>x. ereal (f x)"]
-  using positive_integral_positive[of "\<lambda>x. ereal (-f x)"]
+  using positive_integral_positive[of M "\<lambda>x. ereal (f x)"]
+  using positive_integral_positive[of M "\<lambda>x. ereal (-f x)"]
   by (cases rule: ereal2_cases[of "(\<integral>\<^isup>+x. ereal (-f x)\<partial>M)" "(\<integral>\<^isup>+x. ereal (f x)\<partial>M)"]) auto
 
-lemma (in measure_space) integral_cong:
+lemma integral_cong:
   assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
   shows "integral\<^isup>L M f = integral\<^isup>L M g"
   using assms by (simp cong: positive_integral_cong add: lebesgue_integral_def)
 
-lemma (in measure_space) integral_cong_measure:
-  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
-  shows "integral\<^isup>L N f = integral\<^isup>L M f"
-  by (simp add: positive_integral_cong_measure[OF assms] lebesgue_integral_def)
-
-lemma (in measure_space) integrable_cong_measure:
-  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
-  shows "integrable N f \<longleftrightarrow> integrable M f"
-  using assms
-  by (simp add: positive_integral_cong_measure[OF assms] integrable_def measurable_def)
-
-lemma (in measure_space) integral_cong_AE:
-  assumes cong: "AE x. f x = g x"
+lemma integral_cong_AE:
+  assumes cong: "AE x in M. f x = g x"
   shows "integral\<^isup>L M f = integral\<^isup>L M g"
 proof -
-  have *: "AE x. ereal (f x) = ereal (g x)"
-    "AE x. ereal (- f x) = ereal (- g x)" using cong by auto
+  have *: "AE x in M. ereal (f x) = ereal (g x)"
+    "AE x in M. ereal (- f x) = ereal (- g x)" using cong by auto
   show ?thesis
     unfolding *[THEN positive_integral_cong_AE] lebesgue_integral_def ..
 qed
 
-lemma (in measure_space) integrable_cong_AE:
+lemma integrable_cong_AE:
   assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  assumes "AE x. f x = g x"
+  assumes "AE x in M. f x = g x"
   shows "integrable M f = integrable M g"
 proof -
   have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) = (\<integral>\<^isup>+ x. ereal (g x) \<partial>M)"
@@ -1740,11 +1478,23 @@
     by (auto simp: integrable_def)
 qed
 
-lemma (in measure_space) integrable_cong:
+lemma integrable_cong:
   "(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable M f \<longleftrightarrow> integrable M g"
   by (simp cong: positive_integral_cong measurable_cong add: integrable_def)
 
-lemma (in measure_space) integral_eq_positive_integral:
+lemma positive_integral_eq_integral:
+  assumes f: "integrable M f"
+  assumes nonneg: "AE x in M. 0 \<le> f x" 
+  shows "(\<integral>\<^isup>+ x. ereal (f x) \<partial>M) = integral\<^isup>L M f"
+proof -
+  have "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M) = (\<integral>\<^isup>+ x. 0 \<partial>M)"
+    using nonneg by (intro positive_integral_cong_AE) (auto split: split_max)
+  with f positive_integral_positive show ?thesis
+    by (cases "\<integral>\<^isup>+ x. ereal (f x) \<partial>M")
+       (auto simp add: lebesgue_integral_def positive_integral_max_0 integrable_def)
+qed
+  
+lemma integral_eq_positive_integral:
   assumes f: "\<And>x. 0 \<le> f x"
   shows "integral\<^isup>L M f = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>M)"
 proof -
@@ -1755,84 +1505,12 @@
     unfolding lebesgue_integral_def by simp
 qed
 
-lemma (in measure_space) integral_vimage:
-  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
-  assumes f: "f \<in> borel_measurable M'"
-  shows "integral\<^isup>L M' f = (\<integral>x. f (T x) \<partial>M)"
-proof -
-  interpret T: measure_space M' by (rule measure_space_vimage[OF T])
-  from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
-  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
-    and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
-    using f by (auto simp: comp_def)
-  then show ?thesis
-    using f unfolding lebesgue_integral_def integrable_def
-    by (auto simp: borel[THEN positive_integral_vimage[OF T]])
-qed
-
-lemma (in measure_space) integrable_vimage:
-  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
-  assumes f: "integrable M' f"
-  shows "integrable M (\<lambda>x. f (T x))"
-proof -
-  interpret T: measure_space M' by (rule measure_space_vimage[OF T])
-  from measurable_comp[OF measure_preservingD2[OF T(2)], of f borel]
-  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
-    and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
-    using f by (auto simp: comp_def)
-  then show ?thesis
-    using f unfolding lebesgue_integral_def integrable_def
-    by (auto simp: borel[THEN positive_integral_vimage[OF T]])
-qed
-
-lemma (in measure_space) integral_translated_density:
-  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
-    and g: "g \<in> borel_measurable M"
-    and N: "space N = space M" "sets N = sets M"
-    and density: "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
-      (is "\<And>A. _ \<Longrightarrow> _ = ?d A")
-  shows "integral\<^isup>L N g = (\<integral> x. f x * g x \<partial>M)" (is ?integral)
-    and "integrable N g = integrable M (\<lambda>x. f x * g x)" (is ?integrable)
-proof -
-  from f have ms: "measure_space (M\<lparr>measure := ?d\<rparr>)"
-    by (intro measure_space_density[where u="\<lambda>x. ereal (f x)"]) auto
-
-  from ms density N have "(\<integral>\<^isup>+ x. g x \<partial>N) =  (\<integral>\<^isup>+ x. max 0 (ereal (g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
-    unfolding positive_integral_max_0
-    by (intro measure_space.positive_integral_cong_measure) auto
-  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (f x) * max 0 (ereal (g x)) \<partial>M)"
-    using f g by (intro positive_integral_translated_density) auto
-  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (ereal (f x * g x)) \<partial>M)"
-    using f by (intro positive_integral_cong_AE)
-               (auto simp: ereal_max_0 zero_le_mult_iff split: split_max)
-  finally have pos: "(\<integral>\<^isup>+ x. g x \<partial>N) = (\<integral>\<^isup>+ x. f x * g x \<partial>M)"
-    by (simp add: positive_integral_max_0)
-  
-  from ms density N have "(\<integral>\<^isup>+ x. - (g x) \<partial>N) =  (\<integral>\<^isup>+ x. max 0 (ereal (- g x)) \<partial>M\<lparr>measure := ?d\<rparr>)"
-    unfolding positive_integral_max_0
-    by (intro measure_space.positive_integral_cong_measure) auto
-  also have "\<dots> = (\<integral>\<^isup>+ x. ereal (f x) * max 0 (ereal (- g x)) \<partial>M)"
-    using f g by (intro positive_integral_translated_density) auto
-  also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (ereal (- f x * g x)) \<partial>M)"
-    using f by (intro positive_integral_cong_AE)
-               (auto simp: ereal_max_0 mult_le_0_iff split: split_max)
-  finally have neg: "(\<integral>\<^isup>+ x. - g x \<partial>N) = (\<integral>\<^isup>+ x. - (f x * g x) \<partial>M)"
-    by (simp add: positive_integral_max_0)
-
-  have g_N: "g \<in> borel_measurable N"
-    using g N unfolding measurable_def by simp
-
-  show ?integral ?integrable
-    unfolding lebesgue_integral_def integrable_def
-    using pos neg f g g_N by auto
-qed
-
-lemma (in measure_space) integral_minus[intro, simp]:
+lemma integral_minus[intro, simp]:
   assumes "integrable M f"
   shows "integrable M (\<lambda>x. - f x)" "(\<integral>x. - f x \<partial>M) = - integral\<^isup>L M f"
   using assms by (auto simp: integrable_def lebesgue_integral_def)
 
-lemma (in measure_space) integral_minus_iff[simp]:
+lemma integral_minus_iff[simp]:
   "integrable M (\<lambda>x. - f x) \<longleftrightarrow> integrable M f"
 proof
   assume "integrable M (\<lambda>x. - f x)"
@@ -1841,7 +1519,7 @@
   then show "integrable M f" by simp
 qed (rule integral_minus)
 
-lemma (in measure_space) integral_of_positive_diff:
+lemma integral_of_positive_diff:
   assumes integrable: "integrable M u" "integrable M v"
   and f_def: "\<And>x. f x = u x - v x" and pos: "\<And>x. 0 \<le> u x" "\<And>x. 0 \<le> v x"
   shows "integrable M f" and "integral\<^isup>L M f = integral\<^isup>L M u - integral\<^isup>L M v"
@@ -1851,7 +1529,7 @@
   let ?u = "\<lambda>x. max 0 (ereal (u x))"
   let ?v = "\<lambda>x. max 0 (ereal (v x))"
 
-  from borel_measurable_diff[of u v] integrable
+  from borel_measurable_diff[of u M v] integrable
   have f_borel: "?f \<in> borel_measurable M" and
     mf_borel: "?mf \<in> borel_measurable M" and
     v_borel: "?v \<in> borel_measurable M" and
@@ -1881,7 +1559,7 @@
     using integrable f by (auto elim!: integrableE)
 qed
 
-lemma (in measure_space) integral_linear:
+lemma integral_linear:
   assumes "integrable M f" "integrable M g" and "0 \<le> a"
   shows "integrable M (\<lambda>t. a * f t + g t)"
   and "(\<integral> t. a * f t + g t \<partial>M) = a * integral\<^isup>L M f + integral\<^isup>L M g" (is ?EQ)
@@ -1917,18 +1595,18 @@
     by (auto elim!: integrableE simp: field_simps)
 qed
 
-lemma (in measure_space) integral_add[simp, intro]:
+lemma integral_add[simp, intro]:
   assumes "integrable M f" "integrable M g"
   shows "integrable M (\<lambda>t. f t + g t)"
   and "(\<integral> t. f t + g t \<partial>M) = integral\<^isup>L M f + integral\<^isup>L M g"
   using assms integral_linear[where a=1] by auto
 
-lemma (in measure_space) integral_zero[simp, intro]:
+lemma integral_zero[simp, intro]:
   shows "integrable M (\<lambda>x. 0)" "(\<integral> x.0 \<partial>M) = 0"
   unfolding integrable_def lebesgue_integral_def
   by (auto simp add: borel_measurable_const)
 
-lemma (in measure_space) integral_cmult[simp, intro]:
+lemma integral_cmult[simp, intro]:
   assumes "integrable M f"
   shows "integrable M (\<lambda>t. a * f t)" (is ?P)
   and "(\<integral> t. a * f t \<partial>M) = a * integral\<^isup>L M f" (is ?I)
@@ -1942,37 +1620,80 @@
     assume "a \<le> 0" hence "0 \<le> - a" by auto
     have *: "\<And>t. - a * t + 0 = (-a) * t" by simp
     show ?thesis using integral_linear[OF assms integral_zero(1) `0 \<le> - a`]
-        integral_minus(1)[of "\<lambda>t. - a * f t"]
+        integral_minus(1)[of M "\<lambda>t. - a * f t"]
       unfolding * integral_zero by simp
   qed
   thus ?P ?I by auto
 qed
 
-lemma (in measure_space) integral_multc:
+lemma lebesgue_integral_cmult_nonneg:
+  assumes f: "f \<in> borel_measurable M" and "0 \<le> c"
+  shows "(\<integral>x. c * f x \<partial>M) = c * integral\<^isup>L M f"
+proof -
+  { have "c * real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (f x)))) =
+      real (ereal c * integral\<^isup>P M (\<lambda>x. max 0 (ereal (f x))))"
+      by simp
+    also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. ereal c * max 0 (ereal (f x))))"
+      using f `0 \<le> c` by (subst positive_integral_cmult) auto
+    also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (c * f x))))"
+      using `0 \<le> c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def zero_le_mult_iff)
+    finally have "real (integral\<^isup>P M (\<lambda>x. ereal (c * f x))) = c * real (integral\<^isup>P M (\<lambda>x. ereal (f x)))"
+      by (simp add: positive_integral_max_0) }
+  moreover
+  { have "c * real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (- f x)))) =
+      real (ereal c * integral\<^isup>P M (\<lambda>x. max 0 (ereal (- f x))))"
+      by simp
+    also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. ereal c * max 0 (ereal (- f x))))"
+      using f `0 \<le> c` by (subst positive_integral_cmult) auto
+    also have "\<dots> = real (integral\<^isup>P M (\<lambda>x. max 0 (ereal (- c * f x))))"
+      using `0 \<le> c` by (auto intro!: arg_cong[where f=real] positive_integral_cong simp: max_def mult_le_0_iff)
+    finally have "real (integral\<^isup>P M (\<lambda>x. ereal (- c * f x))) = c * real (integral\<^isup>P M (\<lambda>x. ereal (- f x)))"
+      by (simp add: positive_integral_max_0) }
+  ultimately show ?thesis
+    by (simp add: lebesgue_integral_def field_simps)
+qed
+
+lemma lebesgue_integral_uminus:
+  "(\<integral>x. - f x \<partial>M) = - integral\<^isup>L M f"
+    unfolding lebesgue_integral_def by simp
+
+lemma lebesgue_integral_cmult:
+  assumes f: "f \<in> borel_measurable M"
+  shows "(\<integral>x. c * f x \<partial>M) = c * integral\<^isup>L M f"
+proof (cases rule: linorder_le_cases)
+  assume "0 \<le> c" with f show ?thesis by (rule lebesgue_integral_cmult_nonneg)
+next
+  assume "c \<le> 0"
+  with lebesgue_integral_cmult_nonneg[OF f, of "-c"]
+  show ?thesis
+    by (simp add: lebesgue_integral_def)
+qed
+
+lemma integral_multc:
   assumes "integrable M f"
   shows "(\<integral> x. f x * c \<partial>M) = integral\<^isup>L M f * c"
   unfolding mult_commute[of _ c] integral_cmult[OF assms] ..
 
-lemma (in measure_space) integral_mono_AE:
+lemma integral_mono_AE:
   assumes fg: "integrable M f" "integrable M g"
-  and mono: "AE t. f t \<le> g t"
+  and mono: "AE t in M. f t \<le> g t"
   shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
 proof -
-  have "AE x. ereal (f x) \<le> ereal (g x)"
+  have "AE x in M. ereal (f x) \<le> ereal (g x)"
     using mono by auto
-  moreover have "AE x. ereal (- g x) \<le> ereal (- f x)"
+  moreover have "AE x in M. ereal (- g x) \<le> ereal (- f x)"
     using mono by auto
   ultimately show ?thesis using fg
     by (auto intro!: add_mono positive_integral_mono_AE real_of_ereal_positive_mono
              simp: positive_integral_positive lebesgue_integral_def diff_minus)
 qed
 
-lemma (in measure_space) integral_mono:
+lemma integral_mono:
   assumes "integrable M f" "integrable M g" "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"
   shows "integral\<^isup>L M f \<le> integral\<^isup>L M g"
   using assms by (auto intro: integral_mono_AE)
 
-lemma (in measure_space) integral_diff[simp, intro]:
+lemma integral_diff[simp, intro]:
   assumes f: "integrable M f" and g: "integrable M g"
   shows "integrable M (\<lambda>t. f t - g t)"
   and "(\<integral> t. f t - g t \<partial>M) = integral\<^isup>L M f - integral\<^isup>L M g"
@@ -1980,9 +1701,9 @@
   unfolding diff_minus integral_minus(2)[OF g]
   by auto
 
-lemma (in measure_space) integral_indicator[simp, intro]:
-  assumes "A \<in> sets M" and "\<mu> A \<noteq> \<infinity>"
-  shows "integral\<^isup>L M (indicator A) = real (\<mu> A)" (is ?int)
+lemma integral_indicator[simp, intro]:
+  assumes "A \<in> sets M" and "(emeasure M) A \<noteq> \<infinity>"
+  shows "integral\<^isup>L M (indicator A) = real ((emeasure M) A)" (is ?int)
   and "integrable M (indicator A)" (is ?able)
 proof -
   from `A \<in> sets M` have *:
@@ -1994,10 +1715,10 @@
     by (auto simp: * positive_integral_indicator borel_measurable_indicator)
 qed
 
-lemma (in measure_space) integral_cmul_indicator:
-  assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<infinity>"
+lemma integral_cmul_indicator:
+  assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> (emeasure M) A \<noteq> \<infinity>"
   shows "integrable M (\<lambda>x. c * indicator A x)" (is ?P)
-  and "(\<integral>x. c * indicator A x \<partial>M) = c * real (\<mu> A)" (is ?I)
+  and "(\<integral>x. c * indicator A x \<partial>M) = c * real ((emeasure M) A)" (is ?I)
 proof -
   show ?P
   proof (cases "c = 0")
@@ -2010,7 +1731,7 @@
   qed simp
 qed
 
-lemma (in measure_space) integral_setsum[simp, intro]:
+lemma integral_setsum[simp, intro]:
   assumes "\<And>n. n \<in> S \<Longrightarrow> integrable M (f n)"
   shows "(\<integral>x. (\<Sum> i \<in> S. f i x) \<partial>M) = (\<Sum> i \<in> S. integral\<^isup>L M (f i))" (is "?int S")
     and "integrable M (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I S")
@@ -2023,7 +1744,7 @@
   thus "?int S" and "?I S" by auto
 qed
 
-lemma (in measure_space) integrable_abs:
+lemma integrable_abs:
   assumes "integrable M f"
   shows "integrable M (\<lambda> x. \<bar>f x\<bar>)"
 proof -
@@ -2034,23 +1755,22 @@
     by (simp add: positive_integral_add positive_integral_max_0 integrable_def)
 qed
 
-lemma (in measure_space) integral_subalgebra:
+lemma integral_subalgebra:
   assumes borel: "f \<in> borel_measurable N"
-  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A" and sa: "sigma_algebra N"
+  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A"
   shows "integrable N f \<longleftrightarrow> integrable M f" (is ?P)
     and "integral\<^isup>L N f = integral\<^isup>L M f" (is ?I)
 proof -
-  interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
   have "(\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>M)"
        "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>N) = (\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M)"
-    using borel by (auto intro!: positive_integral_subalgebra N sa)
+    using borel by (auto intro!: positive_integral_subalgebra N)
   moreover have "f \<in> borel_measurable M \<longleftrightarrow> f \<in> borel_measurable N"
     using assms unfolding measurable_def by auto
   ultimately show ?P ?I
     by (auto simp: integrable_def lebesgue_integral_def positive_integral_max_0)
 qed
 
-lemma (in measure_space) integrable_bound:
+lemma integrable_bound:
   assumes "integrable M f"
   and f: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
     "\<And>x. x \<in> space M \<Longrightarrow> \<bar>g x\<bar> \<le> f x"
@@ -2077,11 +1797,21 @@
     unfolding integrable_def by auto
 qed
 
-lemma (in measure_space) integrable_abs_iff:
+lemma lebesgue_integral_nonneg:
+  assumes ae: "(AE x in M. 0 \<le> f x)" shows "0 \<le> integral\<^isup>L M f"
+proof -
+  have "(\<integral>\<^isup>+x. max 0 (ereal (- f x)) \<partial>M) = (\<integral>\<^isup>+x. 0 \<partial>M)"
+    using ae by (intro positive_integral_cong_AE) (auto simp: max_def)
+  then show ?thesis
+    by (auto simp: lebesgue_integral_def positive_integral_max_0
+             intro!: real_of_ereal_pos positive_integral_positive)
+qed
+
+lemma integrable_abs_iff:
   "f \<in> borel_measurable M \<Longrightarrow> integrable M (\<lambda> x. \<bar>f x\<bar>) \<longleftrightarrow> integrable M f"
   by (auto intro!: integrable_bound[where g=f] integrable_abs)
 
-lemma (in measure_space) integrable_max:
+lemma integrable_max:
   assumes int: "integrable M f" "integrable M g"
   shows "integrable M (\<lambda> x. max (f x) (g x))"
 proof (rule integrable_bound)
@@ -2095,7 +1825,7 @@
     by auto
 qed
 
-lemma (in measure_space) integrable_min:
+lemma integrable_min:
   assumes int: "integrable M f" "integrable M g"
   shows "integrable M (\<lambda> x. min (f x) (g x))"
 proof (rule integrable_bound)
@@ -2109,18 +1839,18 @@
     by auto
 qed
 
-lemma (in measure_space) integral_triangle_inequality:
+lemma integral_triangle_inequality:
   assumes "integrable M f"
   shows "\<bar>integral\<^isup>L M f\<bar> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
 proof -
   have "\<bar>integral\<^isup>L M f\<bar> = max (integral\<^isup>L M f) (- integral\<^isup>L M f)" by auto
   also have "\<dots> \<le> (\<integral>x. \<bar>f x\<bar> \<partial>M)"
-      using assms integral_minus(2)[of f, symmetric]
+      using assms integral_minus(2)[of M f, symmetric]
       by (auto intro!: integral_mono integrable_abs simp del: integral_minus)
   finally show ?thesis .
 qed
 
-lemma (in measure_space) integral_positive:
+lemma integral_positive:
   assumes "integrable M f" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
   shows "0 \<le> integral\<^isup>L M f"
 proof -
@@ -2130,7 +1860,7 @@
   finally show ?thesis .
 qed
 
-lemma (in measure_space) integral_monotone_convergence_pos:
+lemma integral_monotone_convergence_pos:
   assumes i: "\<And>i. integrable M (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)"
   and pos: "\<And>x i. 0 \<le> f i x"
   and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
@@ -2157,7 +1887,7 @@
     using i borel_u pos pos_u by (auto simp: positive_integral_0_iff_AE integrable_def)
 
   have integral_eq: "\<And>n. (\<integral>\<^isup>+ x. ereal (f n x) \<partial>M) = ereal (integral\<^isup>L M (f n))"
-    using i positive_integral_positive by (auto simp: ereal_real lebesgue_integral_def integrable_def)
+    using i positive_integral_positive[of M] by (auto simp: ereal_real lebesgue_integral_def integrable_def)
 
   have pos_integral: "\<And>n. 0 \<le> integral\<^isup>L M (f n)"
     using pos i by (auto simp: integral_positive)
@@ -2177,7 +1907,7 @@
     unfolding integrable_def lebesgue_integral_def by auto
 qed
 
-lemma (in measure_space) integral_monotone_convergence:
+lemma integral_monotone_convergence:
   assumes f: "\<And>i. integrable M (f i)" and "mono f"
   and lim: "\<And>x. (\<lambda>i. f i x) ----> u x"
   and ilim: "(\<lambda>i. integral\<^isup>L M (f i)) ----> x"
@@ -2201,9 +1931,9 @@
     by (auto simp: integral_diff)
 qed
 
-lemma (in measure_space) integral_0_iff:
+lemma integral_0_iff:
   assumes "integrable M f"
-  shows "(\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0"
+  shows "(\<integral>x. \<bar>f x\<bar> \<partial>M) = 0 \<longleftrightarrow> (emeasure M) {x\<in>space M. f x \<noteq> 0} = 0"
 proof -
   have *: "(\<integral>\<^isup>+x. ereal (- \<bar>f x\<bar>) \<partial>M) = 0"
     using assms by (auto simp: positive_integral_0_iff_AE integrable_def)
@@ -2212,57 +1942,57 @@
     "(\<integral>\<^isup>+ x. ereal \<bar>f x\<bar> \<partial>M) \<noteq> \<infinity>" unfolding integrable_def by auto
   from positive_integral_0_iff[OF this(1)] this(2)
   show ?thesis unfolding lebesgue_integral_def *
-    using positive_integral_positive[of "\<lambda>x. ereal \<bar>f x\<bar>"]
+    using positive_integral_positive[of M "\<lambda>x. ereal \<bar>f x\<bar>"]
     by (auto simp add: real_of_ereal_eq_0)
 qed
 
-lemma (in measure_space) positive_integral_PInf:
+lemma positive_integral_PInf:
   assumes f: "f \<in> borel_measurable M"
   and not_Inf: "integral\<^isup>P M f \<noteq> \<infinity>"
-  shows "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
+  shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
 proof -
-  have "\<infinity> * \<mu> (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^isup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
+  have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^isup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
     using f by (subst positive_integral_cmult_indicator) (auto simp: measurable_sets)
   also have "\<dots> \<le> integral\<^isup>P M (\<lambda>x. max 0 (f x))"
     by (auto intro!: positive_integral_mono simp: indicator_def max_def)
-  finally have "\<infinity> * \<mu> (f -` {\<infinity>} \<inter> space M) \<le> integral\<^isup>P M f"
+  finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^isup>P M f"
     by (simp add: positive_integral_max_0)
-  moreover have "0 \<le> \<mu> (f -` {\<infinity>} \<inter> space M)"
-    using f by (simp add: measurable_sets)
+  moreover have "0 \<le> (emeasure M) (f -` {\<infinity>} \<inter> space M)"
+    by (rule emeasure_nonneg)
   ultimately show ?thesis
     using assms by (auto split: split_if_asm)
 qed
 
-lemma (in measure_space) positive_integral_PInf_AE:
-  assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<infinity>" shows "AE x. f x \<noteq> \<infinity>"
+lemma positive_integral_PInf_AE:
+  assumes "f \<in> borel_measurable M" "integral\<^isup>P M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>"
 proof (rule AE_I)
-  show "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
+  show "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
     by (rule positive_integral_PInf[OF assms])
   show "f -` {\<infinity>} \<inter> space M \<in> sets M"
     using assms by (auto intro: borel_measurable_vimage)
 qed auto
 
-lemma (in measure_space) simple_integral_PInf:
+lemma simple_integral_PInf:
   assumes "simple_function M f" "\<And>x. 0 \<le> f x"
   and "integral\<^isup>S M f \<noteq> \<infinity>"
-  shows "\<mu> (f -` {\<infinity>} \<inter> space M) = 0"
+  shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
 proof (rule positive_integral_PInf)
   show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function)
   show "integral\<^isup>P M f \<noteq> \<infinity>"
     using assms by (simp add: positive_integral_eq_simple_integral)
 qed
 
-lemma (in measure_space) integral_real:
-  "AE x. \<bar>f x\<bar> \<noteq> \<infinity> \<Longrightarrow> (\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\<lambda>x. - f x))"
+lemma integral_real:
+  "AE x in M. \<bar>f x\<bar> \<noteq> \<infinity> \<Longrightarrow> (\<integral>x. real (f x) \<partial>M) = real (integral\<^isup>P M f) - real (integral\<^isup>P M (\<lambda>x. - f x))"
   using assms unfolding lebesgue_integral_def
   by (subst (1 2) positive_integral_cong_AE) (auto simp add: ereal_real)
 
 lemma (in finite_measure) lebesgue_integral_const[simp]:
   shows "integrable M (\<lambda>x. a)"
-  and  "(\<integral>x. a \<partial>M) = a * \<mu>' (space M)"
+  and  "(\<integral>x. a \<partial>M) = a * (measure M) (space M)"
 proof -
   { fix a :: real assume "0 \<le> a"
-    then have "(\<integral>\<^isup>+ x. ereal a \<partial>M) = ereal a * \<mu> (space M)"
+    then have "(\<integral>\<^isup>+ x. ereal a \<partial>M) = ereal a * (emeasure M) (space M)"
       by (subst positive_integral_const) auto
     moreover
     from `0 \<le> a` have "(\<integral>\<^isup>+ x. ereal (-a) \<partial>M) = 0"
@@ -2277,8 +2007,8 @@
     then have "0 \<le> -a" by auto
     from *[OF this] show ?thesis by simp
   qed
-  show "(\<integral>x. a \<partial>M) = a * \<mu>' (space M)"
-    by (simp add: \<mu>'_def lebesgue_integral_def positive_integral_const_If)
+  show "(\<integral>x. a \<partial>M) = a * measure M (space M)"
+    by (simp add: lebesgue_integral_def positive_integral_const_If emeasure_eq_measure)
 qed
 
 lemma indicator_less[simp]:
@@ -2287,8 +2017,8 @@
 
 lemma (in finite_measure) integral_less_AE:
   assumes int: "integrable M X" "integrable M Y"
-  assumes A: "\<mu> A \<noteq> 0" "A \<in> sets M" "AE x. x \<in> A \<longrightarrow> X x \<noteq> Y x"
-  assumes gt: "AE x. X x \<le> Y x"
+  assumes A: "(emeasure M) A \<noteq> 0" "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> X x \<noteq> Y x"
+  assumes gt: "AE x in M. X x \<le> Y x"
   shows "integral\<^isup>L M X < integral\<^isup>L M Y"
 proof -
   have "integral\<^isup>L M X \<le> integral\<^isup>L M Y"
@@ -2301,26 +2031,27 @@
       using gt by (intro integral_cong_AE) auto
     also have "\<dots> = 0"
       using eq int by simp
-    finally have "\<mu> {x \<in> space M. Y x - X x \<noteq> 0} = 0"
+    finally have "(emeasure M) {x \<in> space M. Y x - X x \<noteq> 0} = 0"
       using int by (simp add: integral_0_iff)
     moreover
     have "(\<integral>\<^isup>+x. indicator A x \<partial>M) \<le> (\<integral>\<^isup>+x. indicator {x \<in> space M. Y x - X x \<noteq> 0} x \<partial>M)"
       using A by (intro positive_integral_mono_AE) auto
-    then have "\<mu> A \<le> \<mu> {x \<in> space M. Y x - X x \<noteq> 0}"
+    then have "(emeasure M) A \<le> (emeasure M) {x \<in> space M. Y x - X x \<noteq> 0}"
       using int A by (simp add: integrable_def)
-    moreover note `\<mu> A \<noteq> 0` positive_measure[OF `A \<in> sets M`]
-    ultimately show False by auto
+    ultimately have "emeasure M A = 0"
+      using emeasure_nonneg[of M A] by simp
+    with `(emeasure M) A \<noteq> 0` show False by auto
   qed
   ultimately show ?thesis by auto
 qed
 
 lemma (in finite_measure) integral_less_AE_space:
   assumes int: "integrable M X" "integrable M Y"
-  assumes gt: "AE x. X x < Y x" "\<mu> (space M) \<noteq> 0"
+  assumes gt: "AE x in M. X x < Y x" "(emeasure M) (space M) \<noteq> 0"
   shows "integral\<^isup>L M X < integral\<^isup>L M Y"
   using gt by (intro integral_less_AE[OF int, where A="space M"]) auto
 
-lemma (in measure_space) integral_dominated_convergence:
+lemma integral_dominated_convergence:
   assumes u: "\<And>i. integrable M (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
   and w: "integrable M w"
   and u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
@@ -2336,7 +2067,7 @@
 
   from u[unfolded integrable_def]
   have u'_borel: "u' \<in> borel_measurable M"
-    using u' by (blast intro: borel_measurable_LIMSEQ[of u])
+    using u' by (blast intro: borel_measurable_LIMSEQ[of M u])
 
   { fix x assume x: "x \<in> space M"
     then have "0 \<le> \<bar>u 0 x\<bar>" by auto
@@ -2385,7 +2116,7 @@
         using diff_less_2w[of _ n] unfolding positive_integral_max_0
         by (intro positive_integral_mono) auto
       then have "?f n = 0"
-        using positive_integral_positive[of ?f'] eq_0 by auto }
+        using positive_integral_positive[of M ?f'] eq_0 by auto }
     then show ?thesis by (simp add: Limsup_const)
   next
     assume neq_0: "(\<integral>\<^isup>+ x. max 0 (ereal (2 * w x)) \<partial>M) \<noteq> 0" (is "?wx \<noteq> 0")
@@ -2411,10 +2142,10 @@
     also have "\<dots> = (\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M) -
         limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"
       unfolding PI_diff positive_integral_max_0
-      using positive_integral_positive[of "\<lambda>x. ereal (2 * w x)"]
+      using positive_integral_positive[of M "\<lambda>x. ereal (2 * w x)"]
       by (subst liminf_ereal_cminus) auto
     finally show ?thesis
-      using neq_0 I2w_fin positive_integral_positive[of "\<lambda>x. ereal (2 * w x)"] pos
+      using neq_0 I2w_fin positive_integral_positive[of M "\<lambda>x. ereal (2 * w x)"] pos
       unfolding positive_integral_max_0
       by (cases rule: ereal2_cases[of "\<integral>\<^isup>+ x. ereal (2 * w x) \<partial>M" "limsup (\<lambda>n. \<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M)"])
          auto
@@ -2430,8 +2161,8 @@
   ultimately have liminf_limsup_eq: "liminf ?f = ereal 0" "limsup ?f = ereal 0"
     using `limsup ?f = 0` by auto
   have "\<And>n. (\<integral>\<^isup>+ x. ereal \<bar>u n x - u' x\<bar> \<partial>M) = ereal (\<integral>x. \<bar>u n x - u' x\<bar> \<partial>M)"
-    using diff positive_integral_positive
-    by (subst integral_eq_positive_integral) (auto simp: ereal_real integrable_def)
+    using diff positive_integral_positive[of M]
+    by (subst integral_eq_positive_integral[of _ M]) (auto simp: ereal_real integrable_def)
   then show ?lim_diff
     using ereal_Liminf_eq_Limsup[OF trivial_limit_sequentially liminf_limsup_eq]
     by (simp add: lim_ereal)
@@ -2456,7 +2187,7 @@
   qed
 qed
 
-lemma (in measure_space) integral_sums:
+lemma integral_sums:
   assumes borel: "\<And>i. integrable M (f i)"
   and summable: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>f i x\<bar>)"
   and sums: "summable (\<lambda>i. (\<integral>x. \<bar>f i x\<bar> \<partial>M))"
@@ -2513,18 +2244,18 @@
 
 section "Lebesgue integration on countable spaces"
 
-lemma (in measure_space) integral_on_countable:
+lemma integral_on_countable:
   assumes f: "f \<in> borel_measurable M"
   and bij: "bij_betw enum S (f ` space M)"
   and enum_zero: "enum ` (-S) \<subseteq> {0}"
-  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<infinity>"
-  and abs_summable: "summable (\<lambda>r. \<bar>enum r * real (\<mu> (f -` {enum r} \<inter> space M))\<bar>)"
+  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> (emeasure M) (f -` {x} \<inter> space M) \<noteq> \<infinity>"
+  and abs_summable: "summable (\<lambda>r. \<bar>enum r * real ((emeasure M) (f -` {enum r} \<inter> space M))\<bar>)"
   shows "integrable M f"
-  and "(\<lambda>r. enum r * real (\<mu> (f -` {enum r} \<inter> space M))) sums integral\<^isup>L M f" (is ?sums)
+  and "(\<lambda>r. enum r * real ((emeasure M) (f -` {enum r} \<inter> space M))) sums integral\<^isup>L M f" (is ?sums)
 proof -
   let ?A = "\<lambda>r. f -` {enum r} \<inter> space M"
   let ?F = "\<lambda>r x. enum r * indicator (?A r) x"
-  have enum_eq: "\<And>r. enum r * real (\<mu> (?A r)) = integral\<^isup>L M (?F r)"
+  have enum_eq: "\<And>r. enum r * real ((emeasure M) (?A r)) = integral\<^isup>L M (?F r)"
     using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
 
   { fix x assume "x \<in> space M"
@@ -2551,9 +2282,9 @@
   { fix r
     have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = (\<integral>x. \<bar>enum r\<bar> * indicator (?A r) x \<partial>M)"
       by (auto simp: indicator_def intro!: integral_cong)
-    also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))"
+    also have "\<dots> = \<bar>enum r\<bar> * real ((emeasure M) (?A r))"
       using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator)
-    finally have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = \<bar>enum r * real (\<mu> (?A r))\<bar>"
+    finally have "(\<integral>x. \<bar>?F r x\<bar> \<partial>M) = \<bar>enum r * real ((emeasure M) (?A r))\<bar>"
       using f by (subst (2) abs_mult_pos[symmetric]) (auto intro!: real_of_ereal_pos measurable_sets) }
   note int_abs_F = this
 
@@ -2570,67 +2301,469 @@
   show "integrable M f" unfolding int_f by simp
 qed
 
-section "Lebesgue integration on finite space"
+section {* Distributions *}
+
+lemma simple_function_distr[simp]:
+  "simple_function (distr M M' T) f \<longleftrightarrow> simple_function M' (\<lambda>x. f x)"
+ unfolding simple_function_def by simp
+
+lemma positive_integral_distr:
+  assumes T: "T \<in> measurable M M'"
+  and f: "f \<in> borel_measurable M'"
+  shows "integral\<^isup>P (distr M M' T) f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
+proof -
+  from borel_measurable_implies_simple_function_sequence'[OF f]
+  guess f' . note f' = this
+  then have f_distr: "\<And>i. simple_function (distr M M' T) (f' i)"
+    by simp
+  let ?f = "\<lambda>i x. f' i (T x)"
+  have inc: "incseq ?f" using f' by (force simp: le_fun_def incseq_def)
+  have sup: "\<And>x. (SUP i. ?f i x) = max 0 (f (T x))"
+    using f'(4) .
+  have sf: "\<And>i. simple_function M (\<lambda>x. f' i (T x))"
+    using simple_function_comp[OF T(1) f'(1)] .
+  show "integral\<^isup>P (distr M M' T) f = (\<integral>\<^isup>+ x. f (T x) \<partial>M)"
+    using
+      positive_integral_monotone_convergence_simple[OF f'(2,5) f_distr]
+      positive_integral_monotone_convergence_simple[OF inc f'(5) sf]
+    by (simp add: positive_integral_max_0 simple_integral_distr[OF T f'(1)] f')
+qed
+
+lemma integral_distr:
+  assumes T: "T \<in> measurable M M'"
+  assumes f: "f \<in> borel_measurable M'"
+  shows "integral\<^isup>L (distr M M' T) f = (\<integral>x. f (T x) \<partial>M)"
+proof -
+  from measurable_comp[OF T, of f borel]
+  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
+    and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
+    using f by (auto simp: comp_def)
+  then show ?thesis
+    using f unfolding lebesgue_integral_def integrable_def
+    by (auto simp: borel[THEN positive_integral_distr[OF T]])
+qed
 
-lemma (in measure_space) integral_on_finite:
-  assumes f: "f \<in> borel_measurable M" and finite: "finite (f`space M)"
-  and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<infinity>"
-  shows "integrable M f"
-  and "(\<integral>x. f x \<partial>M) =
-    (\<Sum> r \<in> f`space M. r * real (\<mu> (f -` {r} \<inter> space M)))" (is "?integral")
+lemma integrable_distr:
+  assumes T: "T \<in> measurable M M'" and f: "integrable (distr M M' T) f"
+  shows "integrable M (\<lambda>x. f (T x))"
+proof -
+  from measurable_comp[OF T, of f borel]
+  have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable M'" "(\<lambda>x. ereal (- f x)) \<in> borel_measurable M'"
+    and "(\<lambda>x. f (T x)) \<in> borel_measurable M"
+    using f by (auto simp: comp_def)
+  then show ?thesis
+    using f unfolding lebesgue_integral_def integrable_def
+    using borel[THEN positive_integral_distr[OF T]]
+    by (auto simp: borel[THEN positive_integral_distr[OF T]])
+qed
+
+section {* Lebesgue integration on @{const count_space} *}
+
+lemma simple_function_count_space[simp]:
+  "simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)"
+  unfolding simple_function_def by simp
+
+lemma positive_integral_count_space:
+  assumes A: "finite {a\<in>A. 0 < f a}"
+  shows "integral\<^isup>P (count_space A) f = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
 proof -
-  let ?A = "\<lambda>r. f -` {r} \<inter> space M"
-  let ?S = "\<lambda>x. \<Sum>r\<in>f`space M. r * indicator (?A r) x"
+  have *: "(\<integral>\<^isup>+x. max 0 (f x) \<partial>count_space A) =
+    (\<integral>\<^isup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
+    by (auto intro!: positive_integral_cong
+             simp add: indicator_def if_distrib setsum_cases[OF A] max_def le_less)
+  also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^isup>+ x. f a * indicator {a} x \<partial>count_space A)"
+    by (subst positive_integral_setsum)
+       (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le)
+  also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)"
+    by (auto intro!: setsum_cong simp: positive_integral_cmult_indicator one_ereal_def[symmetric])
+  finally show ?thesis by (simp add: positive_integral_max_0)
+qed
+
+lemma integrable_count_space:
+  "finite X \<Longrightarrow> integrable (count_space X) f"
+  by (auto simp: positive_integral_count_space integrable_def)
+
+lemma positive_integral_count_space_finite:
+    "finite A \<Longrightarrow> (\<integral>\<^isup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. max 0 (f a))"
+  by (subst positive_integral_max_0[symmetric])
+     (auto intro!: setsum_mono_zero_left simp: positive_integral_count_space less_le)
+
+lemma lebesgue_integral_count_space_finite:
+    "finite A \<Longrightarrow> (\<integral>x. f x \<partial>count_space A) = (\<Sum>a\<in>A. f a)"
+  apply (auto intro!: setsum_mono_zero_left
+              simp: positive_integral_count_space_finite lebesgue_integral_def)
+  apply (subst (1 2)  setsum_real_of_ereal[symmetric])
+  apply (auto simp: max_def setsum_subtractf[symmetric] intro!: setsum_cong)
+  done
+
+section {* Measure spaces with an associated density *}
+
+definition density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
+  "density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
 
-  { fix x assume "x \<in> space M"
-    have "f x = (\<Sum>r\<in>f`space M. if x \<in> ?A r then r else 0)"
-      using finite `x \<in> space M` by (simp add: setsum_cases)
-    also have "\<dots> = ?S x"
-      by (auto intro!: setsum_cong)
-    finally have "f x = ?S x" . }
-  note f_eq = this
+lemma 
+  shows sets_density[simp]: "sets (density M f) = sets M"
+    and space_density[simp]: "space (density M f) = space M"
+  by (auto simp: density_def)
+
+lemma 
+  shows measurable_density_eq1[simp]: "g \<in> measurable (density Mg f) Mg' \<longleftrightarrow> g \<in> measurable Mg Mg'"
+    and measurable_density_eq2[simp]: "h \<in> measurable Mh (density Mh' f) \<longleftrightarrow> h \<in> measurable Mh Mh'"
+    and simple_function_density_eq[simp]: "simple_function (density Mu f) u \<longleftrightarrow> simple_function Mu u"
+  unfolding measurable_def simple_function_def by simp_all
+
+lemma density_cong: "f \<in> borel_measurable M \<Longrightarrow> f' \<in> borel_measurable M \<Longrightarrow>
+  (AE x in M. f x = f' x) \<Longrightarrow> density M f = density M f'"
+  unfolding density_def by (auto intro!: measure_of_eq positive_integral_cong_AE space_closed)
+
+lemma density_max_0: "density M f = density M (\<lambda>x. max 0 (f x))"
+proof -
+  have "\<And>x A. max 0 (f x) * indicator A x = max 0 (f x * indicator A x)"
+    by (auto simp: indicator_def)
+  then show ?thesis
+    unfolding density_def by (simp add: positive_integral_max_0)
+qed
+
+lemma density_ereal_max_0: "density M (\<lambda>x. ereal (f x)) = density M (\<lambda>x. ereal (max 0 (f x)))"
+  by (subst density_max_0) (auto intro!: arg_cong[where f="density M"] split: split_max)
 
-  have f_eq_S: "integrable M f \<longleftrightarrow> integrable M ?S" "integral\<^isup>L M f = integral\<^isup>L M ?S"
-    by (auto intro!: integrable_cong integral_cong simp only: f_eq)
+lemma emeasure_density:
+  assumes f: "f \<in> borel_measurable M" and A: "A \<in> sets M"
+  shows "emeasure (density M f) A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>M)"
+    (is "_ = ?\<mu> A")
+  unfolding density_def
+proof (rule emeasure_measure_of_sigma)
+  show "sigma_algebra (space M) (sets M)" ..
+  show "positive (sets M) ?\<mu>"
+    using f by (auto simp: positive_def intro!: positive_integral_positive)
+  have \<mu>_eq: "?\<mu> = (\<lambda>A. \<integral>\<^isup>+ x. max 0 (f x) * indicator A x \<partial>M)" (is "?\<mu> = ?\<mu>'")
+    apply (subst positive_integral_max_0[symmetric])
+    apply (intro ext positive_integral_cong_AE AE_I2)
+    apply (auto simp: indicator_def)
+    done
+  show "countably_additive (sets M) ?\<mu>"
+    unfolding \<mu>_eq
+  proof (intro countably_additiveI)
+    fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M"
+    then have *: "\<And>i. (\<lambda>x. max 0 (f x) * indicator (A i) x) \<in> borel_measurable M"
+      using f by (auto intro!: borel_measurable_ereal_times)
+    assume disj: "disjoint_family A"
+    have "(\<Sum>n. ?\<mu>' (A n)) = (\<integral>\<^isup>+ x. (\<Sum>n. max 0 (f x) * indicator (A n) x) \<partial>M)"
+      using f * by (simp add: positive_integral_suminf)
+    also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) * (\<Sum>n. indicator (A n) x) \<partial>M)" using f
+      by (auto intro!: suminf_cmult_ereal positive_integral_cong_AE)
+    also have "\<dots> = (\<integral>\<^isup>+ x. max 0 (f x) * indicator (\<Union>n. A n) x \<partial>M)"
+      unfolding suminf_indicator[OF disj] ..
+    finally show "(\<Sum>n. ?\<mu>' (A n)) = ?\<mu>' (\<Union>x. A x)" by simp
+  qed
+qed fact
 
-  show "integrable M f" ?integral using fin f f_eq_S
-    by (simp_all add: integral_cmul_indicator borel_measurable_vimage)
+lemma null_sets_density_iff:
+  assumes f: "f \<in> borel_measurable M"
+  shows "A \<in> null_sets (density M f) \<longleftrightarrow> A \<in> sets M \<and> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
+proof -
+  { assume "A \<in> sets M"
+    have eq: "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. max 0 (f x) * indicator A x \<partial>M)"
+      apply (subst positive_integral_max_0[symmetric])
+      apply (intro positive_integral_cong)
+      apply (auto simp: indicator_def)
+      done
+    have "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> 
+      emeasure M {x \<in> space M. max 0 (f x) * indicator A x \<noteq> 0} = 0"
+      unfolding eq
+      using f `A \<in> sets M`
+      by (intro positive_integral_0_iff) auto
+    also have "\<dots> \<longleftrightarrow> (AE x in M. max 0 (f x) * indicator A x = 0)"
+      using f `A \<in> sets M`
+      by (intro AE_iff_measurable[OF _ refl, symmetric])
+         (auto intro!: sets_Collect borel_measurable_ereal_eq)
+    also have "(AE x in M. max 0 (f x) * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)"
+      by (auto simp add: indicator_def max_def split: split_if_asm)
+    finally have "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" . }
+  with f show ?thesis
+    by (simp add: null_sets_def emeasure_density cong: conj_cong)
+qed
+
+lemma AE_density:
+  assumes f: "f \<in> borel_measurable M"
+  shows "(AE x in density M f. P x) \<longleftrightarrow> (AE x in M. 0 < f x \<longrightarrow> P x)"
+proof
+  assume "AE x in density M f. P x"
+  with f obtain N where "{x \<in> space M. \<not> P x} \<subseteq> N" "N \<in> sets M" and ae: "AE x in M. x \<in> N \<longrightarrow> f x \<le> 0"
+    by (auto simp: eventually_ae_filter null_sets_density_iff)
+  then have "AE x in M. x \<notin> N \<longrightarrow> P x" by auto
+  with ae show "AE x in M. 0 < f x \<longrightarrow> P x"
+    by (rule eventually_elim2) auto
+next
+  fix N assume ae: "AE x in M. 0 < f x \<longrightarrow> P x"
+  then obtain N where "{x \<in> space M. \<not> (0 < f x \<longrightarrow> P x)} \<subseteq> N" "N \<in> null_sets M"
+    by (auto simp: eventually_ae_filter)
+  then have *: "{x \<in> space (density M f). \<not> P x} \<subseteq> N \<union> {x\<in>space M. \<not> 0 < f x}"
+    "N \<union> {x\<in>space M. \<not> 0 < f x} \<in> sets M" and ae2: "AE x in M. x \<notin> N"
+    using f by (auto simp: subset_eq intro!: sets_Collect_neg AE_not_in)
+  show "AE x in density M f. P x"
+    using ae2
+    unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f]
+    by (intro exI[of _ "N \<union> {x\<in>space M. \<not> 0 < f x}"] conjI *)
+       (auto elim: eventually_elim2)
 qed
 
-lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function M f"
-  unfolding simple_function_def using finite_space by auto
-
-lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \<in> borel_measurable M"
-  by (auto intro: borel_measurable_simple_function)
+lemma positive_integral_density:
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  assumes g': "g' \<in> borel_measurable M"
+  shows "integral\<^isup>P (density M f) g' = (\<integral>\<^isup>+ x. f x * g' x \<partial>M)"
+proof -
+  def g \<equiv> "\<lambda>x. max 0 (g' x)"
+  then have g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
+    using g' by auto
+  from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess G . note G = this
+  note G' = borel_measurable_simple_function[OF this(1)] simple_functionD[OF G(1)]
+  note G'(2)[simp]
+  { fix P have "AE x in M. P x \<Longrightarrow> AE x in M. P x"
+      using positive_integral_null_set[of _ _ f]
+      by (auto simp: eventually_ae_filter ) }
+  note ac = this
+  with G(4) f g have G_M': "AE x in density M f. (SUP i. G i x) = g x"
+    by (auto simp add: AE_density[OF f(1)] max_def)
+  { fix i
+    let ?I = "\<lambda>y x. indicator (G i -` {y} \<inter> space M) x"
+    { fix x assume *: "x \<in> space M" "0 \<le> f x" "0 \<le> g x"
+      then have [simp]: "G i ` space M \<inter> {y. G i x = y \<and> x \<in> space M} = {G i x}" by auto
+      from * G' G have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * (\<Sum>y\<in>G i`space M. (y * ?I y x))"
+        by (subst setsum_ereal_right_distrib) (auto simp: ac_simps)
+      also have "\<dots> = f x * G i x"
+        by (simp add: indicator_def if_distrib setsum_cases)
+      finally have "(\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) = f x * G i x" . }
+    note to_singleton = this
+    have "integral\<^isup>P (density M f) (G i) = integral\<^isup>S (density M f) (G i)"
+      using G by (intro positive_integral_eq_simple_integral) simp_all
+    also have "\<dots> = (\<Sum>y\<in>G i`space M. y * (\<integral>\<^isup>+x. f x * ?I y x \<partial>M))"
+      using f G(1)
+      by (auto intro!: setsum_cong arg_cong2[where f="op *"] emeasure_density
+               simp: simple_function_def simple_integral_def)
+    also have "\<dots> = (\<Sum>y\<in>G i`space M. (\<integral>\<^isup>+x. y * (f x * ?I y x) \<partial>M))"
+      using f G' G by (auto intro!: setsum_cong positive_integral_cmult[symmetric])
+    also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>y\<in>G i`space M. y * (f x * ?I y x)) \<partial>M)"
+      using f G' G by (auto intro!: positive_integral_setsum[symmetric])
+    finally have "integral\<^isup>P (density M f) (G i) = (\<integral>\<^isup>+x. f x * G i x \<partial>M)"
+      using f g G' to_singleton by (auto intro!: positive_integral_cong_AE) }
+  note [simp] = this
+  have "integral\<^isup>P (density M f) g = (SUP i. integral\<^isup>P (density M f) (G i))" using G'(1) G_M'(1) G
+    using positive_integral_monotone_convergence_SUP[symmetric, OF `incseq G`, of "density M f"]
+    by (simp cong: positive_integral_cong_AE)
+  also have "\<dots> = (SUP i. (\<integral>\<^isup>+x. f x * G i x \<partial>M))" by simp
+  also have "\<dots> = (\<integral>\<^isup>+x. (SUP i. f x * G i x) \<partial>M)"
+    using f G' G(2)[THEN incseq_SucD] G
+    by (intro positive_integral_monotone_convergence_SUP_AE[symmetric])
+       (auto simp: ereal_mult_left_mono le_fun_def ereal_zero_le_0_iff)
+  also have "\<dots> = (\<integral>\<^isup>+x. f x * g x \<partial>M)" using f G' G g
+    by (intro positive_integral_cong_AE)
+       (auto simp add: SUPR_ereal_cmult split: split_max)
+  also have "\<dots> = (\<integral>\<^isup>+x. f x * g' x \<partial>M)"
+    using f(2)
+    by (subst (2) positive_integral_max_0[symmetric])
+       (auto simp: g_def max_def ereal_zero_le_0_iff intro!: positive_integral_cong_AE)
+  finally show "integral\<^isup>P (density M f) g' = (\<integral>\<^isup>+x. f x * g' x \<partial>M)"
+    unfolding g_def positive_integral_max_0 .
+qed
 
-lemma (in finite_measure_space) positive_integral_finite_eq_setsum:
-  assumes pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x"
-  shows "integral\<^isup>P M f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
+lemma integral_density:
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+    and g: "g \<in> borel_measurable M"
+  shows "integral\<^isup>L (density M f) g = (\<integral> x. f x * g x \<partial>M)"
+    and "integrable (density M f) g \<longleftrightarrow> integrable M (\<lambda>x. f x * g x)"
+  unfolding lebesgue_integral_def integrable_def using f g
+  by (auto simp: positive_integral_density)
+
+lemma emeasure_restricted:
+  assumes S: "S \<in> sets M" and X: "X \<in> sets M"
+  shows "emeasure (density M (indicator S)) X = emeasure M (S \<inter> X)"
 proof -
-  have *: "integral\<^isup>P M f = (\<integral>\<^isup>+ x. (\<Sum>y\<in>space M. f y * indicator {y} x) \<partial>M)"
-    by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space])
-  show ?thesis unfolding * using borel_measurable_finite[of f] pos
-    by (simp add: positive_integral_setsum positive_integral_cmult_indicator)
+  have "emeasure (density M (indicator S)) X = (\<integral>\<^isup>+x. indicator S x * indicator X x \<partial>M)"
+    using S X by (simp add: emeasure_density)
+  also have "\<dots> = (\<integral>\<^isup>+x. indicator (S \<inter> X) x \<partial>M)"
+    by (auto intro!: positive_integral_cong simp: indicator_def)
+  also have "\<dots> = emeasure M (S \<inter> X)"
+    using S X by (simp add: Int)
+  finally show ?thesis .
+qed
+
+lemma measure_restricted:
+  "S \<in> sets M \<Longrightarrow> X \<in> sets M \<Longrightarrow> measure (density M (indicator S)) X = measure M (S \<inter> X)"
+  by (simp add: emeasure_restricted measure_def)
+
+lemma (in finite_measure) finite_measure_restricted:
+  "S \<in> sets M \<Longrightarrow> finite_measure (density M (indicator S))"
+  by default (simp add: emeasure_restricted)
+
+lemma emeasure_density_const:
+  "A \<in> sets M \<Longrightarrow> 0 \<le> c \<Longrightarrow> emeasure (density M (\<lambda>_. c)) A = c * emeasure M A"
+  by (auto simp: positive_integral_cmult_indicator emeasure_density)
+
+lemma measure_density_const:
+  "A \<in> sets M \<Longrightarrow> 0 < c \<Longrightarrow> c \<noteq> \<infinity> \<Longrightarrow> measure (density M (\<lambda>_. c)) A = real c * measure M A"
+  by (auto simp: emeasure_density_const measure_def)
+
+lemma density_density_eq:
+   "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. 0 \<le> f x \<Longrightarrow>
+   density (density M f) g = density M (\<lambda>x. f x * g x)"
+  by (auto intro!: measure_eqI simp: emeasure_density positive_integral_density ac_simps)
+
+lemma distr_density_distr:
+  assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M"
+    and inv: "\<forall>x\<in>space M. T' (T x) = x"
+  assumes f: "f \<in> borel_measurable M'"
+  shows "distr (density (distr M M' T) f) M T' = density M (f \<circ> T)" (is "?R = ?L")
+proof (rule measure_eqI)
+  fix A assume A: "A \<in> sets ?R"
+  { fix x assume "x \<in> space M"
+    with sets_into_space[OF A]
+    have "indicator (T' -` A \<inter> space M') (T x) = (indicator A x :: ereal)"
+      using T inv by (auto simp: indicator_def measurable_space) }
+  with A T T' f show "emeasure ?R A = emeasure ?L A"
+    by (simp add: measurable_comp emeasure_density emeasure_distr
+                  positive_integral_distr measurable_sets cong: positive_integral_cong)
+qed simp
+
+lemma density_density_divide:
+  fixes f g :: "'a \<Rightarrow> real"
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x"
+  assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0"
+  shows "density (density M f) (\<lambda>x. g x / f x) = density M g"
+proof -
+  have "density M g = density M (\<lambda>x. f x * (g x / f x))"
+    using f g ac by (auto intro!: density_cong measurable_If)
+  then show ?thesis
+    using f g by (subst density_density_eq) auto
 qed
 
-lemma (in finite_measure_space) integral_finite_singleton:
-  shows "integrable M f"
-  and "integral\<^isup>L M f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I)
+section {* Point measure *}
+
+definition point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
+  "point_measure A f = density (count_space A) f"
+
+lemma
+  shows space_point_measure: "space (point_measure A f) = A"
+    and sets_point_measure: "sets (point_measure A f) = Pow A"
+  by (auto simp: point_measure_def)
+
+lemma measurable_point_measure_eq1[simp]:
+  "g \<in> measurable (point_measure A f) M \<longleftrightarrow> g \<in> A \<rightarrow> space M"
+  unfolding point_measure_def by simp
+
+lemma measurable_point_measure_eq2_finite[simp]:
+  "finite A \<Longrightarrow>
+   g \<in> measurable M (point_measure A f) \<longleftrightarrow>
+    (g \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. g -` {a} \<inter> space M \<in> sets M))"
+  unfolding point_measure_def by simp
+
+lemma simple_function_point_measure[simp]:
+  "simple_function (point_measure A f) g \<longleftrightarrow> finite (g ` A)"
+  by (simp add: point_measure_def)
+
+lemma emeasure_point_measure:
+  assumes A: "finite {a\<in>X. 0 < f a}" "X \<subseteq> A"
+  shows "emeasure (point_measure A f) X = (\<Sum>a|a\<in>X \<and> 0 < f a. f a)"
 proof -
-  have *:
-    "(\<integral>\<^isup>+ x. max 0 (ereal (f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (ereal (f x)) * \<mu> {x})"
-    "(\<integral>\<^isup>+ x. max 0 (ereal (- f x)) \<partial>M) = (\<Sum>x \<in> space M. max 0 (ereal (- f x)) * \<mu> {x})"
-    by (simp_all add: positive_integral_finite_eq_setsum)
-  then show "integrable M f" using finite_space finite_measure
-    by (simp add: setsum_Pinfty integrable_def positive_integral_max_0
-             split: split_max)
-  show ?I using finite_measure *
-    apply (simp add: positive_integral_max_0 lebesgue_integral_def)
-    apply (subst (1 2) setsum_real_of_ereal[symmetric])
-    apply (simp_all split: split_max add: setsum_subtractf[symmetric])
-    apply (intro setsum_cong[OF refl])
-    apply (simp split: split_max)
-    done
+  have "{a. (a \<in> X \<longrightarrow> a \<in> A \<and> 0 < f a) \<and> a \<in> X} = {a\<in>X. 0 < f a}"
+    using `X \<subseteq> A` by auto
+  with A show ?thesis
+    by (simp add: emeasure_density positive_integral_count_space ereal_zero_le_0_iff
+                  point_measure_def indicator_def)
 qed
 
+lemma emeasure_point_measure_finite:
+  "finite A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a|a\<in>X. f a)"
+  by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum_mono_zero_left simp: less_le)
+
+lemma null_sets_point_measure_iff:
+  "X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x \<le> 0)"
+ by (auto simp: AE_count_space null_sets_density_iff point_measure_def)
+
+lemma AE_point_measure:
+  "(AE x in point_measure A f. P x) \<longleftrightarrow> (\<forall>x\<in>A. 0 < f x \<longrightarrow> P x)"
+  unfolding point_measure_def
+  by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def)
+
+lemma positive_integral_point_measure:
+  "finite {a\<in>A. 0 < f a \<and> 0 < g a} \<Longrightarrow>
+    integral\<^isup>P (point_measure A f) g = (\<Sum>a|a\<in>A \<and> 0 < f a \<and> 0 < g a. f a * g a)"
+  unfolding point_measure_def
+  apply (subst density_max_0)
+  apply (subst positive_integral_density)
+  apply (simp_all add: AE_count_space positive_integral_density)
+  apply (subst positive_integral_count_space )
+  apply (auto intro!: setsum_cong simp: max_def ereal_zero_less_0_iff)
+  apply (rule finite_subset)
+  prefer 2
+  apply assumption
+  apply auto
+  done
+
+lemma positive_integral_point_measure_finite:
+  "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> g a) \<Longrightarrow>
+    integral\<^isup>P (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
+  by (subst positive_integral_point_measure) (auto intro!: setsum_mono_zero_left simp: less_le)
+
+lemma lebesgue_integral_point_measure_finite:
+  "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow> integral\<^isup>L (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)"
+  by (simp add: lebesgue_integral_count_space_finite AE_count_space integral_density point_measure_def)
+
+lemma integrable_point_measure_finite:
+  "finite A \<Longrightarrow> integrable (point_measure A (\<lambda>x. ereal (f x))) g"
+  unfolding point_measure_def
+  apply (subst density_ereal_max_0)
+  apply (subst integral_density)
+  apply (auto simp: AE_count_space integrable_count_space)
+  done
+
+section {* Uniform measure *}
+
+definition "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)"
+
+lemma
+  shows sets_uniform_measure[simp]: "sets (uniform_measure M A) = sets M"
+    and space_uniform_measure[simp]: "space (uniform_measure M A) = space M"
+  by (auto simp: uniform_measure_def)
+
+lemma emeasure_uniform_measure[simp]:
+  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
+  shows "emeasure (uniform_measure M A) B = emeasure M (A \<inter> B) / emeasure M A"
+proof -
+  from A B have "emeasure (uniform_measure M A) B = (\<integral>\<^isup>+x. (1 / emeasure M A) * indicator (A \<inter> B) x \<partial>M)"
+    by (auto simp add: uniform_measure_def emeasure_density split: split_indicator
+             intro!: positive_integral_cong)
+  also have "\<dots> = emeasure M (A \<inter> B) / emeasure M A"
+    using A B
+    by (subst positive_integral_cmult_indicator) (simp_all add: Int emeasure_nonneg)
+  finally show ?thesis .
+qed
+
+lemma emeasure_neq_0_sets: "emeasure M A \<noteq> 0 \<Longrightarrow> A \<in> sets M"
+  using emeasure_notin_sets[of A M] by blast
+
+lemma measure_uniform_measure[simp]:
+  assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>" and B: "B \<in> sets M"
+  shows "measure (uniform_measure M A) B = measure M (A \<inter> B) / measure M A"
+  using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A
+  by (cases "emeasure M A" "emeasure M (A \<inter> B)" rule: ereal2_cases) (simp_all add: measure_def)
+
+section {* Uniform count measure *}
+
+definition "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)"
+ 
+lemma 
+  shows space_uniform_count_measure: "space (uniform_count_measure A) = A"
+    and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A"
+    unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure)
+ 
+lemma emeasure_uniform_count_measure:
+  "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (uniform_count_measure A) X = card X / card A"
+  by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def)
+ 
+lemma measure_uniform_count_measure:
+  "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> measure (uniform_count_measure A) X = card X / card A"
+  by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def measure_def)
+
 end
--- a/src/HOL/Probability/Lebesgue_Measure.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Lebesgue_Measure.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -9,6 +9,15 @@
   imports Finite_Product_Measure
 begin
 
+lemma borel_measurable_sets:
+  assumes "f \<in> measurable borel M" "A \<in> sets M"
+  shows "f -` A \<in> sets borel"
+  using measurable_sets[OF assms] by simp
+
+lemma measurable_identity[intro,simp]:
+  "(\<lambda>x. x) \<in> measurable M M"
+  unfolding measurable_def by auto
+
 subsection {* Standard Cubes *}
 
 definition cube :: "nat \<Rightarrow> 'a::ordered_euclidean_space set" where
@@ -52,17 +61,13 @@
 
 subsection {* Lebesgue measure *} 
 
-definition lebesgue :: "'a::ordered_euclidean_space measure_space" where
-  "lebesgue = \<lparr> space = UNIV,
-    sets = {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n},
-    measure = \<lambda>A. SUP n. ereal (integral (cube n) (indicator A)) \<rparr>"
+definition lebesgue :: "'a::ordered_euclidean_space measure" where
+  "lebesgue = measure_of UNIV {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n}
+    (\<lambda>A. SUP n. ereal (integral (cube n) (indicator A)))"
 
 lemma space_lebesgue[simp]: "space lebesgue = UNIV"
   unfolding lebesgue_def by simp
 
-lemma lebesgueD: "A \<in> sets lebesgue \<Longrightarrow> (indicator A :: _ \<Rightarrow> real) integrable_on cube n"
-  unfolding lebesgue_def by simp
-
 lemma lebesgueI: "(\<And>n. (indicator A :: _ \<Rightarrow> real) integrable_on cube n) \<Longrightarrow> A \<in> sets lebesgue"
   unfolding lebesgue_def by simp
 
@@ -86,23 +91,23 @@
   "A \<inter> B = {} \<Longrightarrow> (indicator A x::_::monoid_add) + indicator B x = indicator (A \<union> B) x"
   unfolding indicator_def by auto
 
-interpretation lebesgue: sigma_algebra lebesgue
-proof (intro sigma_algebra_iff2[THEN iffD2] conjI allI ballI impI lebesgueI)
-  fix A n assume A: "A \<in> sets lebesgue"
-  have "indicator (space lebesgue - A) = (\<lambda>x. 1 - indicator A x :: real)"
+lemma sigma_algebra_lebesgue:
+  defines "leb \<equiv> {A. \<forall>n. (indicator A :: 'a::ordered_euclidean_space \<Rightarrow> real) integrable_on cube n}"
+  shows "sigma_algebra UNIV leb"
+proof (safe intro!: sigma_algebra_iff2[THEN iffD2])
+  fix A assume A: "A \<in> leb"
+  moreover have "indicator (UNIV - A) = (\<lambda>x. 1 - indicator A x :: real)"
     by (auto simp: fun_eq_iff indicator_def)
-  then show "(indicator (space lebesgue - A) :: _ \<Rightarrow> real) integrable_on cube n"
-    using A by (auto intro!: integrable_sub dest: lebesgueD simp: cube_def)
+  ultimately show "UNIV - A \<in> leb"
+    using A by (auto intro!: integrable_sub simp: cube_def leb_def)
 next
-  fix n show "(indicator {} :: _\<Rightarrow>real) integrable_on cube n"
-    by (auto simp: cube_def indicator_def [abs_def])
+  fix n show "{} \<in> leb"
+    by (auto simp: cube_def indicator_def[abs_def] leb_def)
 next
-  fix A :: "nat \<Rightarrow> 'a set" and n ::nat assume "range A \<subseteq> sets lebesgue"
-  then have A: "\<And>i. (indicator (A i) :: _ \<Rightarrow> real) integrable_on cube n"
-    by (auto dest: lebesgueD)
-  show "(indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n" (is "?g integrable_on _")
-  proof (intro dominated_convergence[where g="?g"] ballI)
-    fix k show "(indicator (\<Union>i<k. A i) :: _ \<Rightarrow> real) integrable_on cube n"
+  fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> leb"
+  have "\<forall>n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n" (is "\<forall>n. ?g integrable_on _")
+  proof (intro dominated_convergence[where g="?g"] ballI allI)
+    fix k n show "(indicator (\<Union>i<k. A i) :: _ \<Rightarrow> real) integrable_on cube n"
     proof (induct k)
       case (Suc k)
       have *: "(\<Union> i<Suc k. A i) = (\<Union> i<k. A i) \<union> A k"
@@ -111,36 +116,45 @@
           indicator (\<Union> i<Suc k. A i)" (is "(\<lambda>x. max (?f x) (?g x)) = _")
         by (auto simp: fun_eq_iff * indicator_def)
       show ?case
-        using absolutely_integrable_max[of ?f "cube n" ?g] A Suc by (simp add: *)
+        using absolutely_integrable_max[of ?f "cube n" ?g] A Suc
+        by (simp add: * leb_def subset_eq)
     qed auto
   qed (auto intro: LIMSEQ_indicator_UN simp: cube_def)
+  then show "(\<Union>i. A i) \<in> leb" by (auto simp: leb_def)
 qed simp
 
-interpretation lebesgue: measure_space lebesgue
-proof
+lemma sets_lebesgue: "sets lebesgue = {A. \<forall>n. (indicator A :: _ \<Rightarrow> real) integrable_on cube n}"
+  unfolding lebesgue_def sigma_algebra.sets_measure_of_eq[OF sigma_algebra_lebesgue] ..
+
+lemma lebesgueD: "A \<in> sets lebesgue \<Longrightarrow> (indicator A :: _ \<Rightarrow> real) integrable_on cube n"
+  unfolding sets_lebesgue by simp
+
+lemma emeasure_lebesgue: 
+  assumes "A \<in> sets lebesgue"
+  shows "emeasure lebesgue A = (SUP n. ereal (integral (cube n) (indicator A)))"
+    (is "_ = ?\<mu> A")
+proof (rule emeasure_measure_of[OF lebesgue_def])
   have *: "indicator {} = (\<lambda>x. 0 :: real)" by (simp add: fun_eq_iff)
-  show "positive lebesgue (measure lebesgue)"
-  proof (unfold positive_def, safe)
-    show "measure lebesgue {} = 0" by (simp add: integral_0 * lebesgue_def)
-    fix A assume "A \<in> sets lebesgue"
-    then show "0 \<le> measure lebesgue A"
-      unfolding lebesgue_def
-      by (auto intro!: SUP_upper2 integral_nonneg)
+  show "positive (sets lebesgue) ?\<mu>"
+  proof (unfold positive_def, intro conjI ballI)
+    show "?\<mu> {} = 0" by (simp add: integral_0 *)
+    fix A :: "'a set" assume "A \<in> sets lebesgue" then show "0 \<le> ?\<mu> A"
+      by (auto intro!: SUP_upper2 Integration.integral_nonneg simp: sets_lebesgue)
   qed
 next
-  show "countably_additive lebesgue (measure lebesgue)"
+  show "countably_additive (sets lebesgue) ?\<mu>"
   proof (intro countably_additive_def[THEN iffD2] allI impI)
-    fix A :: "nat \<Rightarrow> 'b set" assume rA: "range A \<subseteq> sets lebesgue" "disjoint_family A"
+    fix A :: "nat \<Rightarrow> 'a set" assume rA: "range A \<subseteq> sets lebesgue" "disjoint_family A"
     then have A[simp, intro]: "\<And>i n. (indicator (A i) :: _ \<Rightarrow> real) integrable_on cube n"
       by (auto dest: lebesgueD)
     let ?m = "\<lambda>n i. integral (cube n) (indicator (A i) :: _\<Rightarrow>real)"
     let ?M = "\<lambda>n I. integral (cube n) (indicator (\<Union>i\<in>I. A i) :: _\<Rightarrow>real)"
-    have nn[simp, intro]: "\<And>i n. 0 \<le> ?m n i" by (auto intro!: integral_nonneg)
+    have nn[simp, intro]: "\<And>i n. 0 \<le> ?m n i" by (auto intro!: Integration.integral_nonneg)
     assume "(\<Union>i. A i) \<in> sets lebesgue"
     then have UN_A[simp, intro]: "\<And>i n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n"
-      by (auto dest: lebesgueD)
-    show "(\<Sum>n. measure lebesgue (A n)) = measure lebesgue (\<Union>i. A i)"
-    proof (simp add: lebesgue_def, subst suminf_SUP_eq, safe intro!: incseq_SucI)
+      by (auto simp: sets_lebesgue)
+    show "(\<Sum>n. ?\<mu> (A n)) = ?\<mu> (\<Union>i. A i)"
+    proof (subst suminf_SUP_eq, safe intro!: incseq_SucI)
       fix i n show "ereal (?m n i) \<le> ereal (?m (Suc n) i)"
         using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI)
     next
@@ -172,14 +186,15 @@
               indicator (\<Union>i<m. A i) x + (indicator (A m) x :: real)"
               by (auto simp: indicator_add lessThan_Suc ac_simps)
             ultimately show ?case
-              using Suc A by (simp add: integral_add[symmetric])
+              using Suc A by (simp add: Integration.integral_add[symmetric])
           qed auto }
         ultimately show "(\<lambda>m. \<Sum>x = 0..<m. ?m n x) ----> ?M n UNIV"
           by (simp add: atLeast0LessThan)
       qed
     qed
   qed
-qed
+next
+qed (auto, fact)
 
 lemma has_integral_interval_cube:
   fixes a b :: "'a::ordered_euclidean_space"
@@ -202,9 +217,10 @@
   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"
-  proof (safe intro!: lebesgueI)
+  have "s \<in> sigma_sets (space lebesgue) (range (\<lambda>(a, b). {a .. (b :: 'a\<Colon>ordered_euclidean_space)}))"
+    using assms by (simp add: borel_eq_atLeastAtMost)
+  also have "\<dots> \<subseteq> sets lebesgue"
+  proof (safe intro!: sigma_sets_subset lebesgueI)
     fix n :: nat and a b :: 'a
     let ?N = "\<chi>\<chi> i. max (- real n) (a $$ i)"
     let ?P = "\<chi>\<chi> i. min (real n) (b $$ i)"
@@ -212,11 +228,7 @@
       unfolding integrable_on_def
       using has_integral_interval_cube[of a b] by auto
   qed
-  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)
+  finally show ?thesis .
 qed
 
 lemma lebesgueI_negligible[dest]: fixes s::"'a::ordered_euclidean_space set"
@@ -224,19 +236,21 @@
   using assms by (force simp: cube_def integrable_on_def negligible_def intro!: lebesgueI)
 
 lemma lmeasure_eq_0:
-  fixes S :: "'a::ordered_euclidean_space set" assumes "negligible S" shows "lebesgue.\<mu> S = 0"
+  fixes S :: "'a::ordered_euclidean_space set"
+  assumes "negligible S" shows "emeasure lebesgue S = 0"
 proof -
   have "\<And>n. integral (cube n) (indicator S :: 'a\<Rightarrow>real) = 0"
     unfolding lebesgue_integral_def using assms
     by (intro integral_unique some1_equality ex_ex1I)
        (auto simp: cube_def negligible_def)
-  then show ?thesis by (auto simp: lebesgue_def)
+  then show ?thesis
+    using assms by (simp add: emeasure_lebesgue lebesgueI_negligible)
 qed
 
 lemma lmeasure_iff_LIMSEQ:
-  assumes "A \<in> sets lebesgue" "0 \<le> m"
-  shows "lebesgue.\<mu> A = ereal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
-proof (simp add: lebesgue_def, intro SUP_eq_LIMSEQ)
+  assumes A: "A \<in> sets lebesgue" and "0 \<le> m"
+  shows "emeasure lebesgue A = ereal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m"
+proof (subst emeasure_lebesgue[OF A], intro SUP_eq_LIMSEQ)
   show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))"
     using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD)
 qed
@@ -261,7 +275,7 @@
 
 lemma lmeasure_finite_has_integral:
   fixes s :: "'a::ordered_euclidean_space set"
-  assumes "s \<in> sets lebesgue" "lebesgue.\<mu> s = ereal m" "0 \<le> m"
+  assumes "s \<in> sets lebesgue" "emeasure lebesgue s = ereal m" "0 \<le> m"
   shows "(indicator s has_integral m) UNIV"
 proof -
   let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
@@ -275,7 +289,7 @@
            (auto dest!: lebesgueD) }
     moreover
     { fix n have "0 \<le> integral (cube n) (?I s)"
-      using assms by (auto dest!: lebesgueD intro!: integral_nonneg) }
+      using assms by (auto dest!: lebesgueD intro!: Integration.integral_nonneg) }
     ultimately
     show "bounded {integral UNIV (?I (s \<inter> cube k)) |k. True}"
       unfolding bounded_def
@@ -303,14 +317,13 @@
     unfolding m by (intro integrable_integral **)
 qed
 
-lemma lmeasure_finite_integrable: assumes s: "s \<in> sets lebesgue" and "lebesgue.\<mu> s \<noteq> \<infinity>"
+lemma lmeasure_finite_integrable: assumes s: "s \<in> sets lebesgue" and "emeasure lebesgue s \<noteq> \<infinity>"
   shows "(indicator s :: _ \<Rightarrow> real) integrable_on UNIV"
-proof (cases "lebesgue.\<mu> s")
+proof (cases "emeasure lebesgue s")
   case (real m)
-  with lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this]
-    lebesgue.positive_measure[OF s]
+  with lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this] emeasure_nonneg[of lebesgue s]
   show ?thesis unfolding integrable_on_def by auto
-qed (insert assms lebesgue.positive_measure[OF s], auto)
+qed (insert assms emeasure_nonneg[of lebesgue s], auto)
 
 lemma has_integral_lebesgue: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
   shows "s \<in> sets lebesgue"
@@ -324,7 +337,7 @@
 qed
 
 lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV"
-  shows "lebesgue.\<mu> s = ereal m"
+  shows "emeasure lebesgue s = ereal m"
 proof (intro lmeasure_iff_LIMSEQ[THEN iffD2])
   let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real"
   show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] .
@@ -349,55 +362,56 @@
 qed
 
 lemma has_integral_iff_lmeasure:
-  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = ereal m)"
+  "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> emeasure lebesgue A = ereal m)"
 proof
   assume "(indicator A has_integral m) UNIV"
   with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this]
-  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = ereal m"
+  show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> emeasure lebesgue A = ereal m"
     by (auto intro: has_integral_nonneg)
 next
-  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> lebesgue.\<mu> A = ereal m"
+  assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> emeasure lebesgue A = ereal m"
   then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto
 qed
 
 lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV"
-  shows "lebesgue.\<mu> s = ereal (integral UNIV (indicator s))"
+  shows "emeasure lebesgue s = ereal (integral UNIV (indicator s))"
   using assms unfolding integrable_on_def
 proof safe
   fix y :: real assume "(indicator s has_integral y) UNIV"
   from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this]
-  show "lebesgue.\<mu> s = ereal (integral UNIV (indicator s))" by simp
+  show "emeasure lebesgue s = ereal (integral UNIV (indicator s))" by simp
 qed
 
 lemma lebesgue_simple_function_indicator:
   fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
   assumes f:"simple_function lebesgue f"
   shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
-  by (rule, subst lebesgue.simple_function_indicator_representation[OF f]) auto
+  by (rule, subst simple_function_indicator_representation[OF f]) auto
 
 lemma integral_eq_lmeasure:
-  "(indicator s::_\<Rightarrow>real) integrable_on UNIV \<Longrightarrow> integral UNIV (indicator s) = real (lebesgue.\<mu> s)"
+  "(indicator s::_\<Rightarrow>real) integrable_on UNIV \<Longrightarrow> integral UNIV (indicator s) = real (emeasure lebesgue s)"
   by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg)
 
-lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "lebesgue.\<mu> s \<noteq> \<infinity>"
+lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "emeasure lebesgue s \<noteq> \<infinity>"
   using lmeasure_eq_integral[OF assms] by auto
 
 lemma negligible_iff_lebesgue_null_sets:
-  "negligible A \<longleftrightarrow> A \<in> lebesgue.null_sets"
+  "negligible A \<longleftrightarrow> A \<in> null_sets lebesgue"
 proof
   assume "negligible A"
   from this[THEN lebesgueI_negligible] this[THEN lmeasure_eq_0]
-  show "A \<in> lebesgue.null_sets" by auto
+  show "A \<in> null_sets lebesgue" by auto
 next
-  assume A: "A \<in> lebesgue.null_sets"
-  then have *:"((indicator A) has_integral (0::real)) UNIV" using lmeasure_finite_has_integral[of A] by auto
+  assume A: "A \<in> null_sets lebesgue"
+  then have *:"((indicator A) has_integral (0::real)) UNIV" using lmeasure_finite_has_integral[of A]
+    by (auto simp: null_sets_def)
   show "negligible A" unfolding negligible_def
   proof (intro allI)
     fix a b :: 'a
     have integrable: "(indicator A :: _\<Rightarrow>real) integrable_on {a..b}"
       by (intro integrable_on_subinterval has_integral_integrable) (auto intro: *)
     then have "integral {a..b} (indicator A) \<le> (integral UNIV (indicator A) :: real)"
-      using * by (auto intro!: integral_subset_le has_integral_integrable)
+      using * by (auto intro!: integral_subset_le)
     moreover have "(0::real) \<le> integral {a..b} (indicator A)"
       using integrable by (auto intro!: integral_nonneg)
     ultimately have "integral {a..b} (indicator A) = (0::real)"
@@ -412,8 +426,8 @@
   shows "integral {a .. b} (\<lambda>x. c) = content {a .. b} *\<^sub>R c"
   by (rule integral_unique) (rule has_integral_const)
 
-lemma lmeasure_UNIV[intro]: "lebesgue.\<mu> (UNIV::'a::ordered_euclidean_space set) = \<infinity>"
-proof (simp add: lebesgue_def, intro SUP_PInfty bexI)
+lemma lmeasure_UNIV[intro]: "emeasure lebesgue (UNIV::'a::ordered_euclidean_space set) = \<infinity>"
+proof (simp add: emeasure_lebesgue, intro SUP_PInfty bexI)
   fix n :: nat
   have "indicator UNIV = (\<lambda>x::'a. 1 :: real)" by auto
   moreover
@@ -434,7 +448,7 @@
 
 lemma
   fixes a b ::"'a::ordered_euclidean_space"
-  shows lmeasure_atLeastAtMost[simp]: "lebesgue.\<mu> {a..b} = ereal (content {a..b})"
+  shows lmeasure_atLeastAtMost[simp]: "emeasure lebesgue {a..b} = ereal (content {a..b})"
 proof -
   have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV"
     unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def])
@@ -454,7 +468,7 @@
 qed
 
 lemma lmeasure_singleton[simp]:
-  fixes a :: "'a::ordered_euclidean_space" shows "lebesgue.\<mu> {a} = 0"
+  fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0"
   using lmeasure_atLeastAtMost[of a a] by simp
 
 declare content_real[simp]
@@ -462,82 +476,68 @@
 lemma
   fixes a b :: real
   shows lmeasure_real_greaterThanAtMost[simp]:
-    "lebesgue.\<mu> {a <.. b} = ereal (if a \<le> b then b - a else 0)"
+    "emeasure lebesgue {a <.. b} = ereal (if a \<le> b then b - a else 0)"
 proof cases
   assume "a < b"
-  then have "lebesgue.\<mu> {a <.. b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {a}"
-    by (subst lebesgue.measure_Diff[symmetric])
-       (auto intro!: arg_cong[where f=lebesgue.\<mu>])
+  then have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b} - emeasure lebesgue {a}"
+    by (subst emeasure_Diff[symmetric])
+       (auto intro!: arg_cong[where f="emeasure lebesgue"])
   then show ?thesis by auto
 qed auto
 
 lemma
   fixes a b :: real
   shows lmeasure_real_atLeastLessThan[simp]:
-    "lebesgue.\<mu> {a ..< b} = ereal (if a \<le> b then b - a else 0)"
+    "emeasure lebesgue {a ..< b} = ereal (if a \<le> b then b - a else 0)"
 proof cases
   assume "a < b"
-  then have "lebesgue.\<mu> {a ..< b} = lebesgue.\<mu> {a .. b} - lebesgue.\<mu> {b}"
-    by (subst lebesgue.measure_Diff[symmetric])
-       (auto intro!: arg_cong[where f=lebesgue.\<mu>])
+  then have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b} - emeasure lebesgue {b}"
+    by (subst emeasure_Diff[symmetric])
+       (auto intro!: arg_cong[where f="emeasure lebesgue"])
   then show ?thesis by auto
 qed auto
 
 lemma
   fixes a b :: real
   shows lmeasure_real_greaterThanLessThan[simp]:
-    "lebesgue.\<mu> {a <..< b} = ereal (if a \<le> b then b - a else 0)"
+    "emeasure lebesgue {a <..< b} = ereal (if a \<le> b then b - a else 0)"
 proof cases
   assume "a < b"
-  then have "lebesgue.\<mu> {a <..< b} = lebesgue.\<mu> {a <.. b} - lebesgue.\<mu> {b}"
-    by (subst lebesgue.measure_Diff[symmetric])
-       (auto intro!: arg_cong[where f=lebesgue.\<mu>])
+  then have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a <.. b} - emeasure lebesgue {b}"
+    by (subst emeasure_Diff[symmetric])
+       (auto intro!: arg_cong[where f="emeasure lebesgue"])
   then show ?thesis by auto
 qed auto
 
 subsection {* Lebesgue-Borel measure *}
 
-definition "lborel = lebesgue \<lparr> sets := sets borel \<rparr>"
+definition "lborel = measure_of UNIV (sets borel) (emeasure lebesgue)"
 
 lemma
   shows space_lborel[simp]: "space lborel = UNIV"
   and sets_lborel[simp]: "sets lborel = sets borel"
-  and measure_lborel[simp]: "measure lborel = lebesgue.\<mu>"
-  and measurable_lborel[simp]: "measurable lborel = measurable borel"
-  by (simp_all add: measurable_def [abs_def] lborel_def)
+  and measurable_lborel1[simp]: "measurable lborel = measurable borel"
+  and measurable_lborel2[simp]: "measurable A lborel = measurable A borel"
+  using sigma_sets_eq[of borel]
+  by (auto simp add: lborel_def measurable_def[abs_def])
 
-interpretation lborel: measure_space "lborel :: ('a::ordered_euclidean_space) measure_space"
-  where "space lborel = UNIV"
-  and "sets lborel = sets borel"
-  and "measure lborel = lebesgue.\<mu>"
-  and "measurable lborel = measurable borel"
-proof (rule lebesgue.measure_space_subalgebra)
-  have "sigma_algebra (lborel::'a measure_space) \<longleftrightarrow> sigma_algebra (borel::'a algebra)"
-    unfolding sigma_algebra_iff2 lborel_def by simp
-  then show "sigma_algebra (lborel::'a measure_space)" by simp default
-qed auto
+lemma emeasure_lborel[simp]: "A \<in> sets borel \<Longrightarrow> emeasure lborel A = emeasure lebesgue A"
+  by (rule emeasure_measure_of[OF lborel_def])
+     (auto simp: positive_def emeasure_nonneg countably_additive_def intro!: suminf_emeasure)
 
 interpretation lborel: sigma_finite_measure lborel
-  where "space lborel = UNIV"
-  and "sets lborel = sets borel"
-  and "measure lborel = lebesgue.\<mu>"
-  and "measurable lborel = measurable borel"
-proof -
-  show "sigma_finite_measure lborel"
-  proof (default, intro conjI exI[of _ "\<lambda>n. cube n"])
-    show "range cube \<subseteq> sets lborel" 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 lborel" by auto
-    show "\<forall>i. measure lborel (cube i) \<noteq> \<infinity>" by (simp add: cube_def)
-  qed
-qed simp_all
+proof (default, intro conjI exI[of _ "\<lambda>n. cube n"])
+  show "range cube \<subseteq> sets lborel" by (auto intro: borel_closed)
+  { fix x :: 'a have "\<exists>n. x\<in>cube n" using mem_big_cube by auto }
+  then show "(\<Union>i. cube i) = (space lborel :: 'a set)" using mem_big_cube by auto
+  show "\<forall>i. emeasure lborel (cube i) \<noteq> \<infinity>" by (simp add: cube_def)
+qed
 
 interpretation lebesgue: sigma_finite_measure lebesgue
 proof
-  from lborel.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. lebesgue.\<mu> (A i) \<noteq> \<infinity>)"
-    by auto
+  from lborel.sigma_finite guess A :: "nat \<Rightarrow> 'a set" ..
+  then show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. emeasure lebesgue (A i) \<noteq> \<infinity>)"
+    by (intro exI[of _ A]) (auto simp: subset_eq)
 qed
 
 subsection {* Lebesgue integrable implies Gauge integrable *}
@@ -556,11 +556,11 @@
   fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"
   assumes f:"simple_function lebesgue f"
   and f':"range f \<subseteq> {0..<\<infinity>}"
-  and om:"\<And>x. x \<in> range f \<Longrightarrow> lebesgue.\<mu> (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> x = 0"
+  and om:"\<And>x. x \<in> range f \<Longrightarrow> emeasure lebesgue (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> x = 0"
   shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
   unfolding simple_integral_def space_lebesgue
 proof (subst lebesgue_simple_function_indicator)
-  let ?M = "\<lambda>x. lebesgue.\<mu> (f -` {x} \<inter> UNIV)"
+  let ?M = "\<lambda>x. emeasure lebesgue (f -` {x} \<inter> UNIV)"
   let ?F = "\<lambda>x. indicator (f -` {x})"
   { fix x y assume "y \<in> range f"
     from subsetD[OF f' this] have "y * ?F y x = ereal (real y * ?F y x)"
@@ -571,7 +571,7 @@
     have "x * ?M x = real x * real (?M x)"
     proof cases
       assume "x \<noteq> 0" with om[OF x] have "?M x \<noteq> \<infinity>" by auto
-      with subsetD[OF f' x] f[THEN lebesgue.simple_functionD(2)] show ?thesis
+      with subsetD[OF f' x] f[THEN simple_functionD(2)] show ?thesis
         by (cases rule: ereal2_cases[of x "?M x"]) auto
     qed simp }
   ultimately
@@ -580,11 +580,11 @@
     by simp
   also have \<dots>
   proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral
-               real_of_ereal_pos lebesgue.positive_measure ballI)
-    show *: "finite (range f)" "\<And>y. f -` {y} \<in> sets lebesgue" "\<And>y. f -` {y} \<inter> UNIV \<in> sets lebesgue"
-      using lebesgue.simple_functionD[OF f] by auto
+               real_of_ereal_pos emeasure_nonneg ballI)
+    show *: "finite (range f)" "\<And>y. f -` {y} \<in> sets lebesgue"
+      using simple_functionD[OF f] by auto
     fix y assume "real y \<noteq> 0" "y \<in> range f"
-    with * om[OF this(2)] show "lebesgue.\<mu> (f -` {y}) = ereal (real (?M y))"
+    with * om[OF this(2)] show "emeasure lebesgue (f -` {y}) = ereal (real (?M y))"
       by (auto simp: ereal_real)
   qed
   finally show "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV" .
@@ -601,28 +601,28 @@
   shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"
 proof -
   let ?f = "\<lambda>x. if x \<in> f -` {\<infinity>} then 0 else f x"
-  note f(1)[THEN lebesgue.simple_functionD(2)]
+  note f(1)[THEN simple_functionD(2)]
   then have [simp, intro]: "\<And>X. f -` X \<in> sets lebesgue" by auto
   have f': "simple_function lebesgue ?f"
-    using f by (intro lebesgue.simple_function_If_set) auto
+    using f by (intro simple_function_If_set) auto
   have rng: "range ?f \<subseteq> {0..<\<infinity>}" using f by auto
   have "AE x in lebesgue. f x = ?f x"
-    using lebesgue.simple_integral_PInf[OF f i]
-    by (intro lebesgue.AE_I[where N="f -` {\<infinity>} \<inter> space lebesgue"]) auto
+    using simple_integral_PInf[OF f i]
+    by (intro AE_I[where N="f -` {\<infinity>} \<inter> space lebesgue"]) auto
   from f(1) f' this have eq: "integral\<^isup>S lebesgue f = integral\<^isup>S lebesgue ?f"
-    by (rule lebesgue.simple_integral_cong_AE)
+    by (rule simple_integral_cong_AE)
   have real_eq: "\<And>x. real (f x) = real (?f x)" by auto
 
   show ?thesis
     unfolding eq real_eq
   proof (rule simple_function_has_integral[OF f' rng])
-    fix x assume x: "x \<in> range ?f" and inf: "lebesgue.\<mu> (?f -` {x} \<inter> UNIV) = \<infinity>"
-    have "x * lebesgue.\<mu> (?f -` {x} \<inter> UNIV) = (\<integral>\<^isup>S y. x * indicator (?f -` {x}) y \<partial>lebesgue)"
-      using f'[THEN lebesgue.simple_functionD(2)]
-      by (simp add: lebesgue.simple_integral_cmult_indicator)
+    fix x assume x: "x \<in> range ?f" and inf: "emeasure lebesgue (?f -` {x} \<inter> UNIV) = \<infinity>"
+    have "x * emeasure lebesgue (?f -` {x} \<inter> UNIV) = (\<integral>\<^isup>S y. x * indicator (?f -` {x}) y \<partial>lebesgue)"
+      using f'[THEN simple_functionD(2)]
+      by (simp add: simple_integral_cmult_indicator)
     also have "\<dots> \<le> integral\<^isup>S lebesgue f"
-      using f'[THEN lebesgue.simple_functionD(2)] f
-      by (intro lebesgue.simple_integral_mono lebesgue.simple_function_mult lebesgue.simple_function_indicator)
+      using f'[THEN simple_functionD(2)] f
+      by (intro simple_integral_mono simple_function_mult simple_function_indicator)
          (auto split: split_indicator)
     finally show "x = 0" unfolding inf using i subsetD[OF rng x] by (auto split: split_if_asm)
   qed
@@ -633,16 +633,16 @@
   assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<noteq> \<infinity>"
   shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV"
 proof -
-  from lebesgue.borel_measurable_implies_simple_function_sequence'[OF f(1)]
+  from borel_measurable_implies_simple_function_sequence'[OF f(1)]
   guess u . note u = this
   have SUP_eq: "\<And>x. (SUP i. u i x) = f x"
     using u(4) f(2)[THEN subsetD] by (auto split: split_max)
   let ?u = "\<lambda>i x. real (u i x)"
-  note u_eq = lebesgue.positive_integral_eq_simple_integral[OF u(1,5), symmetric]
+  note u_eq = positive_integral_eq_simple_integral[OF u(1,5), symmetric]
   { fix i
     note u_eq
     also have "integral\<^isup>P lebesgue (u i) \<le> (\<integral>\<^isup>+x. max 0 (f x) \<partial>lebesgue)"
-      by (intro lebesgue.positive_integral_mono) (auto intro: SUP_upper simp: u(4)[symmetric])
+      by (intro positive_integral_mono) (auto intro: SUP_upper simp: u(4)[symmetric])
     finally have "integral\<^isup>S lebesgue (u i) \<noteq> \<infinity>"
       unfolding positive_integral_max_0 using f by auto }
   note u_fin = this
@@ -684,10 +684,10 @@
       also have "\<dots> = real (integral\<^isup>S lebesgue (u k))"
         using u_int[THEN integral_unique] by (simp add: u')
       also have "\<dots> = real (integral\<^isup>P lebesgue (u k))"
-        using lebesgue.positive_integral_eq_simple_integral[OF u(1,5)] by simp
+        using positive_integral_eq_simple_integral[OF u(1,5)] by simp
       also have "\<dots> \<le> real (integral\<^isup>P lebesgue f)" using f
-        by (auto intro!: real_of_ereal_positive_mono lebesgue.positive_integral_positive
-             lebesgue.positive_integral_mono SUP_upper simp: SUP_eq[symmetric])
+        by (auto intro!: real_of_ereal_positive_mono positive_integral_positive
+             positive_integral_mono SUP_upper simp: SUP_eq[symmetric])
       finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^isup>P lebesgue f)" .
     qed
   qed
@@ -695,21 +695,21 @@
   have "integral\<^isup>P lebesgue f = ereal (integral UNIV f')"
   proof (rule tendsto_unique[OF trivial_limit_sequentially])
     have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))"
-      unfolding u_eq by (intro LIMSEQ_ereal_SUPR lebesgue.incseq_positive_integral u)
-    also note lebesgue.positive_integral_monotone_convergence_SUP
-      [OF u(2)  lebesgue.borel_measurable_simple_function[OF u(1)] u(5), symmetric]
+      unfolding u_eq by (intro LIMSEQ_ereal_SUPR incseq_positive_integral u)
+    also note positive_integral_monotone_convergence_SUP
+      [OF u(2)  borel_measurable_simple_function[OF u(1)] u(5), symmetric]
     finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f"
       unfolding SUP_eq .
 
     { fix k
       have "0 \<le> integral\<^isup>S lebesgue (u k)"
-        using u by (auto intro!: lebesgue.simple_integral_positive)
+        using u by (auto intro!: simple_integral_positive)
       then have "integral\<^isup>S lebesgue (u k) = ereal (real (integral\<^isup>S lebesgue (u k)))"
         using u_fin by (auto simp: ereal_real) }
     note * = this
     show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> ereal (integral UNIV f')"
       using convergent using u_int[THEN integral_unique, symmetric]
-      by (subst *) (simp add: lim_ereal u')
+      by (subst *) (simp add: u')
   qed
   then show ?thesis using convergent by (simp add: f' integrable_integral)
 qed
@@ -721,8 +721,8 @@
 proof -
   let ?n = "\<lambda>x. real (ereal (max 0 (- f x)))" and ?p = "\<lambda>x. real (ereal (max 0 (f x)))"
   have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: ereal_max)
-  { fix f have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (max 0 (f x)) \<partial>lebesgue)"
-      by (intro lebesgue.positive_integral_cong_pos) (auto split: split_max) }
+  { fix f :: "'a \<Rightarrow> real" have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (max 0 (f x)) \<partial>lebesgue)"
+      by (intro positive_integral_cong_pos) (auto split: split_max) }
   note eq = this
   show ?thesis
     unfolding lebesgue_integral_def
@@ -732,7 +732,7 @@
     apply (safe intro!: positive_integral_has_integral)
     using integrableD[OF f]
     by (auto simp: zero_ereal_def[symmetric] positive_integral_max_0  split: split_max
-             intro!: lebesgue.measurable_If lebesgue.borel_measurable_ereal)
+             intro!: measurable_If)
 qed
 
 lemma lebesgue_positive_integral_eq_borel:
@@ -740,7 +740,7 @@
   shows "integral\<^isup>P lebesgue f = integral\<^isup>P lborel f"
 proof -
   from f have "integral\<^isup>P lebesgue (\<lambda>x. max 0 (f x)) = integral\<^isup>P lborel (\<lambda>x. max 0 (f x))"
-    by (auto intro!: lebesgue.positive_integral_subalgebra[symmetric]) default
+    by (auto intro!: positive_integral_subalgebra[symmetric])
   then show ?thesis unfolding positive_integral_max_0 .
 qed
 
@@ -749,9 +749,8 @@
   shows "integrable lebesgue f \<longleftrightarrow> integrable lborel f" (is ?P)
     and "integral\<^isup>L lebesgue f = integral\<^isup>L lborel f" (is ?I)
 proof -
-  have *: "sigma_algebra lborel" by default
   have "sets lborel \<subseteq> sets lebesgue" by auto
-  from lebesgue.integral_subalgebra[of f lborel, OF _ this _ _ *] assms
+  from integral_subalgebra[of f lborel, OF _ this _ _] assms
   show ?P ?I by auto
 qed
 
@@ -783,152 +782,109 @@
   "p2e (e2p x) = (x::'a::ordered_euclidean_space)"
   by (auto simp: euclidean_eq[where 'a='a] p2e_def e2p_def)
 
-interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure_space"
+interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure"
   by default
 
-interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure_space" "{..<n}" for n :: nat
-  where "space lborel = UNIV"
-  and "sets lborel = sets borel"
-  and "measure lborel = lebesgue.\<mu>"
-  and "measurable lborel = measurable borel"
-proof -
-  show "finite_product_sigma_finite (\<lambda>x. lborel::real measure_space) {..<n}"
-    by default simp
-qed simp_all
+interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "{..<n}" for n :: nat
+  by default auto
+
+lemma bchoice_iff: "(\<forall>x\<in>A. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x\<in>A. P x (f x))"
+  by metis
 
 lemma sets_product_borel:
-  assumes [intro]: "finite I"
-  shows "sets (\<Pi>\<^isub>M i\<in>I.
-     \<lparr> space = UNIV::real set, sets = range lessThan, measure = lebesgue.\<mu> \<rparr>) =
-   sets (\<Pi>\<^isub>M i\<in>I. lborel)" (is "sets ?G = _")
-proof -
-  have "sets ?G = sets (\<Pi>\<^isub>M i\<in>I.
-       sigma \<lparr> space = UNIV::real set, sets = range lessThan, measure = lebesgue.\<mu> \<rparr>)"
-    by (subst sigma_product_algebra_sigma_eq[of I "\<lambda>_ i. {..<real i}" ])
-       (auto intro!: measurable_sigma_sigma incseq_SucI reals_Archimedean2
-             simp: product_algebra_def)
-  then show ?thesis
-    unfolding lborel_def borel_eq_lessThan lebesgue_def sigma_def by simp
-qed
+  assumes I: "finite I"
+  shows "sets (\<Pi>\<^isub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^isub>E i\<in>I. UNIV) { \<Pi>\<^isub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G")
+proof (subst sigma_prod_algebra_sigma_eq[where S="\<lambda>_ i::nat. {..<real i}" and E="\<lambda>_. range lessThan", OF I])
+  show "sigma_sets (space (Pi\<^isub>M I (\<lambda>i. lborel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G"
+    by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff)
+qed (auto simp: borel_eq_lessThan incseq_def reals_Archimedean2 image_iff intro: real_natceiling_ge)
 
 lemma measurable_e2p:
-  "e2p \<in> measurable (borel::'a::ordered_euclidean_space algebra)
-                    (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure_space))"
-    (is "_ \<in> measurable ?E ?P")
-proof -
-  let ?B = "\<lparr> space = UNIV::real set, sets = range lessThan, measure = lebesgue.\<mu> \<rparr>"
-  let ?G = "product_algebra_generator {..<DIM('a)} (\<lambda>_. ?B)"
-  have "e2p \<in> measurable ?E (sigma ?G)"
-  proof (rule borel.measurable_sigma)
-    show "e2p \<in> space ?E \<rightarrow> space ?G" by (auto simp: e2p_def)
-    fix A assume "A \<in> sets ?G"
-    then obtain E where A: "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. E i)"
-      and "E \<in> {..<DIM('a)} \<rightarrow> (range lessThan)"
-      by (auto elim!: product_algebraE simp: )
-    then have "\<forall>i\<in>{..<DIM('a)}. \<exists>xs. E i = {..< xs}" by auto
-    from this[THEN bchoice] guess xs ..
-    then have A_eq: "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..< xs i})"
-      using A by auto
-    have "e2p -` A = {..< (\<chi>\<chi> i. xs i) :: 'a}"
-      using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def A_eq
-        euclidean_eq[where 'a='a] eucl_less[where 'a='a])
-    then show "e2p -` A \<inter> space ?E \<in> sets ?E" by simp
-  qed (auto simp: product_algebra_generator_def)
-  with sets_product_borel[of "{..<DIM('a)}"] show ?thesis
-    unfolding measurable_def product_algebra_def by simp
-qed
+  "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))"
+proof (rule measurable_sigma_sets[OF sets_product_borel])
+  fix A :: "(nat \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i} |x. True} "
+  then obtain x where  "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i})" by auto
+  then have "e2p -` A = {..< (\<chi>\<chi> i. x i) :: 'a}"
+    using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def
+      euclidean_eq[where 'a='a] eucl_less[where 'a='a])
+  then show "e2p -` A \<inter> space (borel::'a measure) \<in> sets borel" by simp
+qed (auto simp: e2p_def)
 
 lemma measurable_p2e:
-  "p2e \<in> measurable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure_space))
-    (borel :: 'a::ordered_euclidean_space algebra)"
+  "p2e \<in> measurable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))
+    (borel :: 'a::ordered_euclidean_space measure)"
   (is "p2e \<in> measurable ?P _")
-  unfolding borel_eq_lessThan
-proof (intro lborel_space.measurable_sigma)
-  let ?E = "\<lparr> space = UNIV :: 'a set, sets = range lessThan \<rparr>"
-  show "p2e \<in> space ?P \<rightarrow> space ?E" by simp
-  fix A assume "A \<in> sets ?E"
-  then obtain x where "A = {..<x}" by auto
-  then have "p2e -` A \<inter> space ?P = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..< x $$ i})"
-    using DIM_positive
-    by (auto simp: Pi_iff set_eq_iff p2e_def
-                   euclidean_eq[where 'a='a] eucl_less[where 'a='a])
-  then show "p2e -` A \<inter> space ?P \<in> sets ?P" by auto
-qed simp
+proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2])
+  fix x i
+  let ?A = "{w \<in> space ?P. (p2e w :: 'a) $$ i \<le> x}"
+  assume "i < DIM('a)"
+  then have "?A = (\<Pi>\<^isub>E j\<in>{..<DIM('a)}. if i = j then {.. x} else UNIV)"
+    using DIM_positive by (auto simp: space_PiM p2e_def split: split_if_asm)
+  then show "?A \<in> sets ?P"
+    by auto
+qed
+
+lemma Int_stable_atLeastAtMost:
+  fixes x::"'a::ordered_euclidean_space"
+  shows "Int_stable (range (\<lambda>(a, b::'a). {a..b}))"
+  by (auto simp: inter_interval Int_stable_def)
 
-lemma Int_stable_cuboids:
-  fixes x::"'a::ordered_euclidean_space"
-  shows "Int_stable \<lparr>space = UNIV, sets = range (\<lambda>(a, b::'a). {a..b})\<rparr>"
-  by (auto simp: inter_interval Int_stable_def)
+lemma lborel_eqI:
+  fixes M :: "'a::ordered_euclidean_space measure"
+  assumes emeasure_eq: "\<And>a b. emeasure M {a .. b} = content {a .. b}"
+  assumes sets_eq: "sets M = sets borel"
+  shows "lborel = M"
+proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost])
+  let ?P = "\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel"
+  let ?E = "range (\<lambda>(a, b). {a..b} :: 'a set)"
+  show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"
+    by (simp_all add: borel_eq_atLeastAtMost sets_eq)
+  
+  show "range cube \<subseteq> ?E" unfolding cube_def [abs_def] by auto
+  show "incseq cube" using cube_subset_Suc by (auto intro!: incseq_SucI)
+  { fix x :: 'a have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastforce }
+  then show "(\<Union>i. cube i :: 'a set) = UNIV" by auto
+
+  { fix i show "emeasure lborel (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
+  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
+      by (auto simp: emeasure_eq) }
+qed
 
 lemma lborel_eq_lborel_space:
-  fixes A :: "('a::ordered_euclidean_space) set"
-  assumes "A \<in> sets borel"
-  shows "lborel.\<mu> A = lborel_space.\<mu> DIM('a) (p2e -` A \<inter> (space (lborel_space.P DIM('a))))"
-    (is "_ = measure ?P (?T A)")
-proof (rule measure_unique_Int_stable_vimage)
-  show "measure_space ?P" by default
-  show "measure_space lborel" by default
-
-  let ?E = "\<lparr> space = UNIV :: 'a set, sets = range (\<lambda>(a,b). {a..b}) \<rparr>"
-  show "Int_stable ?E" using Int_stable_cuboids .
-  show "range cube \<subseteq> sets ?E" unfolding cube_def [abs_def] by auto
-  show "incseq cube" using cube_subset_Suc by (auto intro!: incseq_SucI)
-  { fix x have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastforce }
-  then show "(\<Union>i. cube i) = space ?E" by auto
-  { fix i show "lborel.\<mu> (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
-  show "A \<in> sets (sigma ?E)" "sets (sigma ?E) = sets lborel" "space ?E = space lborel"
-    using assms by (simp_all add: borel_eq_atLeastAtMost)
-
-  show "p2e \<in> measurable ?P (lborel :: 'a measure_space)"
-    using measurable_p2e unfolding measurable_def by simp
-  { fix X assume "X \<in> sets ?E"
-    then obtain a b where X[simp]: "X = {a .. b}" by auto
-    have *: "?T X = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {a $$ i .. b $$ i})"
-      by (auto simp: Pi_iff eucl_le[where 'a='a] p2e_def)
-    show "lborel.\<mu> X = measure ?P (?T X)"
-    proof cases
-      assume "X \<noteq> {}"
-      then have "a \<le> b"
-        by (simp add: interval_ne_empty eucl_le[where 'a='a])
-      then have "lborel.\<mu> X = (\<Prod>x<DIM('a). lborel.\<mu> {a $$ x .. b $$ x})"
-        by (auto simp: content_closed_interval eucl_le[where 'a='a]
-                 intro!: setprod_ereal[symmetric])
-      also have "\<dots> = measure ?P (?T X)"
-        unfolding * by (subst lborel_space.measure_times) auto
-      finally show ?thesis .
-    qed simp }
+  "(lborel :: 'a measure) = distr (\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel) lborel p2e"
+  (is "?B = ?D")
+proof (rule lborel_eqI)
+  show "sets ?D = sets borel" by simp
+  let ?P = "(\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel)"
+  fix a b :: 'a
+  have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {a $$ i .. b $$ i})"
+    by (auto simp: Pi_iff eucl_le[where 'a='a] p2e_def space_PiM)
+  have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}"
+  proof cases
+    assume "{a..b} \<noteq> {}"
+    then have "a \<le> b"
+      by (simp add: interval_ne_empty eucl_le[where 'a='a])
+    then have "emeasure lborel {a..b} = (\<Prod>x<DIM('a). emeasure lborel {a $$ x .. b $$ x})"
+      by (auto simp: content_closed_interval eucl_le[where 'a='a]
+               intro!: setprod_ereal[symmetric])
+    also have "\<dots> = emeasure ?P (p2e -` {a..b} \<inter> space ?P)"
+      unfolding * by (subst lborel_space.measure_times) auto
+    finally show ?thesis by simp
+  qed simp
+  then show "emeasure ?D {a .. b} = content {a .. b}"
+    by (simp add: emeasure_distr measurable_p2e)
 qed
 
-lemma measure_preserving_p2e:
-  "p2e \<in> measure_preserving (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure_space))
-    (lborel::'a::ordered_euclidean_space measure_space)" (is "_ \<in> measure_preserving ?P ?E")
-proof
-  show "p2e \<in> measurable ?P ?E"
-    using measurable_p2e by (simp add: measurable_def)
-  fix A :: "'a set" assume "A \<in> sets lborel"
-  then show "lborel_space.\<mu> DIM('a) (p2e -` A \<inter> (space (lborel_space.P DIM('a)))) = lborel.\<mu> A"
-    by (intro lborel_eq_lborel_space[symmetric]) simp
-qed
-
-lemma lebesgue_eq_lborel_space_in_borel:
-  fixes A :: "('a::ordered_euclidean_space) set"
-  assumes A: "A \<in> sets borel"
-  shows "lebesgue.\<mu> A = lborel_space.\<mu> DIM('a) (p2e -` A \<inter> (space (lborel_space.P DIM('a))))"
-  using lborel_eq_lborel_space[OF A] by simp
-
 lemma borel_fubini_positiv_integral:
   fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"
   assumes f: "f \<in> borel_measurable borel"
-  shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(lborel_space.P DIM('a))"
-proof (rule lborel_space.positive_integral_vimage[OF _ measure_preserving_p2e _])
-  show "f \<in> borel_measurable lborel"
-    using f by (simp_all add: measurable_def)
-qed default
+  shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel)"
+  by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f)
 
 lemma borel_fubini_integrable:
   fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
   shows "integrable lborel f \<longleftrightarrow>
-    integrable (lborel_space.P DIM('a)) (\<lambda>x. f (p2e x))"
+    integrable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel) (\<lambda>x. f (p2e x))"
     (is "_ \<longleftrightarrow> integrable ?B ?f")
 proof
   assume "integrable lborel f"
@@ -941,9 +897,9 @@
     by (simp add: comp_def borel_fubini_positiv_integral integrable_def)
 next
   assume "integrable ?B ?f"
-  moreover then
-  have "?f \<circ> e2p \<in> borel_measurable (borel::'a algebra)"
-    by (auto intro!: measurable_e2p measurable_comp)
+  moreover
+  then have "?f \<circ> e2p \<in> borel_measurable (borel::'a measure)"
+    by (auto intro!: measurable_e2p)
   then have "f \<in> borel_measurable borel"
     by (simp cong: measurable_cong)
   ultimately show "integrable lborel f"
@@ -953,100 +909,35 @@
 lemma borel_fubini:
   fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"
   assumes f: "f \<in> borel_measurable borel"
-  shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>(lborel_space.P DIM('a))"
+  shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel))"
   using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def)
 
-
-lemma Int_stable_atLeastAtMost:
-  "Int_stable \<lparr>space = UNIV, sets = range (\<lambda>(a,b). {a::'a::ordered_euclidean_space .. b}) \<rparr>"
-proof (simp add: Int_stable_def image_iff, intro allI)
-  fix a1 b1 a2 b2 :: 'a
-  have "\<forall>i<DIM('a). \<exists>a b. {a1$$i..b1$$i} \<inter> {a2$$i..b2$$i} = {a..b}" by auto
-  then have "\<exists>a b. \<forall>i<DIM('a). {a1$$i..b1$$i} \<inter> {a2$$i..b2$$i} = {a i..b i}"
-    unfolding choice_iff' .
-  then guess a b by safe
-  then have "{a1..b1} \<inter> {a2..b2} = {(\<chi>\<chi> i. a i) .. (\<chi>\<chi> i. b i)}"
-    by (simp add: set_eq_iff eucl_le[where 'a='a] all_conj_distrib[symmetric]) blast
-  then show "\<exists>a' b'. {a1..b1} \<inter> {a2..b2} = {a'..b'}" by blast
-qed
-
-lemma (in sigma_algebra) borel_measurable_sets:
-  assumes "f \<in> measurable borel M" "A \<in> sets M"
-  shows "f -` A \<in> sets borel"
-  using measurable_sets[OF assms] by simp
-
-lemma (in sigma_algebra) measurable_identity[intro,simp]:
-  "(\<lambda>x. x) \<in> measurable M M"
-  unfolding measurable_def by auto
+lemma borel_measurable_indicator':
+  "A \<in> sets borel \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. indicator A (f x)) \<in> borel_measurable M"
+  using measurable_comp[OF _ borel_measurable_indicator, of f M borel A] by (auto simp add: comp_def)
 
 lemma lebesgue_real_affine:
-  fixes X :: "real set"
-  assumes "X \<in> sets borel" and "c \<noteq> 0"
-  shows "measure lebesgue X = ereal \<bar>c\<bar> * measure lebesgue ((\<lambda>x. t + c * x) -` X)"
-    (is "_ = ?\<nu> X")
-proof -
-  let ?E = "\<lparr>space = UNIV, sets = range (\<lambda>(a,b). {a::real .. b})\<rparr> :: real algebra"
-  let ?M = "\<lambda>\<nu>. \<lparr>space = space ?E, sets = sets (sigma ?E), measure = \<nu>\<rparr> :: real measure_space"
-  have *: "?M (measure lebesgue) = lborel"
-    unfolding borel_eq_atLeastAtMost[symmetric]
-    by (simp add: lborel_def lebesgue_def)
-  have **: "?M ?\<nu> = lborel \<lparr> measure := ?\<nu> \<rparr>"
-    unfolding borel_eq_atLeastAtMost[symmetric]
-    by (simp add: lborel_def lebesgue_def)
-  show ?thesis
-  proof (rule measure_unique_Int_stable[where X=X, OF Int_stable_atLeastAtMost], unfold * **)
-    show "X \<in> sets (sigma ?E)"
-      unfolding borel_eq_atLeastAtMost[symmetric] by fact
-    have "\<And>x. \<exists>xa. x \<in> cube xa" apply(rule_tac x=x in mem_big_cube) by fastforce
-    then show "(\<Union>i. cube i) = space ?E" by auto
-    show "incseq cube" by (intro incseq_SucI cube_subset_Suc)
-    show "range cube \<subseteq> sets ?E"
-      unfolding cube_def [abs_def] by auto
-    show "\<And>i. measure lebesgue (cube i) \<noteq> \<infinity>"
-      by (simp add: cube_def)
-    show "measure_space lborel" by default
-    then interpret sigma_algebra "lborel\<lparr>measure := ?\<nu>\<rparr>"
-      by (auto simp add: measure_space_def)
-    show "measure_space (lborel\<lparr>measure := ?\<nu>\<rparr>)"
-    proof
-      show "positive (lborel\<lparr>measure := ?\<nu>\<rparr>) (measure (lborel\<lparr>measure := ?\<nu>\<rparr>))"
-        by (auto simp: positive_def intro!: ereal_0_le_mult borel.borel_measurable_sets)
-      show "countably_additive (lborel\<lparr>measure := ?\<nu>\<rparr>) (measure (lborel\<lparr>measure := ?\<nu>\<rparr>))"
-      proof (simp add: countably_additive_def, safe)
-        fix A :: "nat \<Rightarrow> real set" assume A: "range A \<subseteq> sets borel" "disjoint_family A"
-        then have Ai: "\<And>i. A i \<in> sets borel" by auto
-        have "(\<Sum>n. measure lebesgue ((\<lambda>x. t + c * x) -` A n)) = measure lebesgue (\<Union>n. (\<lambda>x. t + c * x) -` A n)"
-        proof (intro lborel.measure_countably_additive)
-          { fix n have "(\<lambda>x. t + c * x) -` A n \<inter> space borel \<in> sets borel"
-              using A borel.measurable_ident unfolding id_def
-              by (intro measurable_sets[where A=borel] borel.borel_measurable_add[OF _ borel.borel_measurable_times]) auto }
-          then show "range (\<lambda>i. (\<lambda>x. t + c * x) -` A i) \<subseteq> sets borel" by auto
-          from `disjoint_family A`
-          show "disjoint_family (\<lambda>i. (\<lambda>x. t + c * x) -` A i)"
-            by (rule disjoint_family_on_bisimulation) auto
-        qed
-        with Ai show "(\<Sum>n. ?\<nu> (A n)) = ?\<nu> (UNION UNIV A)"
-          by (subst suminf_cmult_ereal)
-             (auto simp: vimage_UN borel.borel_measurable_sets)
-      qed
-    qed
-    fix X assume "X \<in> sets ?E"
-    then obtain a b where [simp]: "X = {a .. b}" by auto
-    show "measure lebesgue X = ?\<nu> X"
-    proof cases
-      assume "0 < c"
-      then have "(\<lambda>x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}"
-        by (auto simp: field_simps)
-      with `0 < c` show ?thesis
-        by (cases "a \<le> b") (auto simp: field_simps)
-    next
-      assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto
-      then have *: "(\<lambda>x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}"
-        by (auto simp: field_simps)
-      with `c < 0` show ?thesis
-        by (cases "a \<le> b") (auto simp: field_simps)
-    qed
+  fixes c :: real assumes "c \<noteq> 0"
+  shows "lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. \<bar>c\<bar>)" (is "_ = ?D")
+proof (rule lborel_eqI)
+  fix a b show "emeasure ?D {a..b} = content {a .. b}"
+  proof cases
+    assume "0 < c"
+    then have "(\<lambda>x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}"
+      by (auto simp: field_simps)
+    with `0 < c` show ?thesis
+      by (cases "a \<le> b")
+         (auto simp: field_simps emeasure_density positive_integral_distr positive_integral_cmult
+                     borel_measurable_indicator' emeasure_distr)
+  next
+    assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto
+    then have *: "(\<lambda>x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}"
+      by (auto simp: field_simps)
+    with `c < 0` show ?thesis
+      by (cases "a \<le> b")
+         (auto simp: field_simps emeasure_density positive_integral_distr
+                     positive_integral_cmult borel_measurable_indicator' emeasure_distr)
   qed
-qed
+qed simp
 
 end
--- a/src/HOL/Probability/Measure.thy	Mon Apr 23 12:23:23 2012 +0100
+++ /dev/null	Thu Jan 01 00:00:00 1970 +0000
@@ -1,1452 +0,0 @@
-(*  Title:      HOL/Probability/Measure.thy
-    Author:     Lawrence C Paulson
-    Author:     Johannes Hölzl, TU München
-    Author:     Armin Heller, TU München
-*)
-
-header {* Properties about measure spaces *}
-
-theory Measure
-  imports Caratheodory
-begin
-
-lemma measure_algebra_more[simp]:
-  "\<lparr> space = A, sets = B, \<dots> = algebra.more M \<rparr> \<lparr> measure := m \<rparr> =
-   \<lparr> space = A, sets = B, \<dots> = algebra.more (M \<lparr> measure := m \<rparr>) \<rparr>"
-  by (cases M) simp
-
-lemma measure_algebra_more_eq[simp]:
-  "\<And>X. measure \<lparr> space = T, sets = A, \<dots> = algebra.more X \<rparr> = measure X"
-  unfolding measure_space.splits by simp
-
-lemma measure_sigma[simp]: "measure (sigma A) = measure A"
-  unfolding sigma_def by simp
-
-lemma algebra_measure_update[simp]:
-  "algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> algebra M'"
-  unfolding algebra_iff_Un by simp
-
-lemma sigma_algebra_measure_update[simp]:
-  "sigma_algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> sigma_algebra M'"
-  unfolding sigma_algebra_def sigma_algebra_axioms_def by simp
-
-lemma finite_sigma_algebra_measure_update[simp]:
-  "finite_sigma_algebra (M'\<lparr>measure := m\<rparr>) \<longleftrightarrow> finite_sigma_algebra M'"
-  unfolding finite_sigma_algebra_def finite_sigma_algebra_axioms_def by simp
-
-lemma measurable_cancel_measure[simp]:
-  "measurable M1 (M2\<lparr>measure := m2\<rparr>) = measurable M1 M2"
-  "measurable (M2\<lparr>measure := m1\<rparr>) M1 = measurable M2 M1"
-  unfolding measurable_def by auto
-
-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>} *}
-
-lemma (in measure_space) measure_countably_additive:
-  assumes "range A \<subseteq> sets M" "disjoint_family A"
-  shows "(\<Sum>i. \<mu> (A i)) = \<mu> (\<Union>i. A i)"
-proof -
-  have "(\<Union> i. A i) \<in> sets M" using assms(1) by (rule countable_UN)
-  with ca assms show ?thesis by (simp add: countably_additive_def)
-qed
-
-lemma (in sigma_algebra) sigma_algebra_cong:
-  assumes "space N = space M" "sets N = sets M"
-  shows "sigma_algebra N"
-  by default (insert sets_into_space, auto simp: assms)
-
-lemma (in measure_space) measure_space_cong:
-  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "space N = space M" "sets N = sets M"
-  shows "measure_space N"
-proof -
-  interpret N: sigma_algebra N by (intro sigma_algebra_cong assms)
-  show ?thesis
-  proof
-    show "positive N (measure N)" using assms by (auto simp: positive_def)
-    show "countably_additive N (measure N)" unfolding countably_additive_def
-    proof safe
-      fix A :: "nat \<Rightarrow> 'a set" assume A: "range A \<subseteq> sets N" "disjoint_family A"
-      then have "\<And>i. A i \<in> sets M" "(UNION UNIV A) \<in> sets M" unfolding assms by auto
-      from measure_countably_additive[of A] A this[THEN assms(1)]
-      show "(\<Sum>n. measure N (A n)) = measure N (UNION UNIV A)"
-        unfolding assms by simp
-    qed
-  qed
-qed
-
-lemma (in measure_space) additive: "additive M \<mu>"
-  using ca by (auto intro!: countably_additive_additive simp: positive_def)
-
-lemma (in measure_space) measure_additive:
-     "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {}
-      \<Longrightarrow> \<mu> a + \<mu> b = \<mu> (a \<union> b)"
-  by (metis additiveD additive)
-
-lemma (in measure_space) measure_mono:
-  assumes "a \<subseteq> b" "a \<in> sets M" "b \<in> sets M"
-  shows "\<mu> a \<le> \<mu> b"
-proof -
-  have "b = a \<union> (b - a)" using assms by auto
-  moreover have "{} = a \<inter> (b - a)" by auto
-  ultimately have "\<mu> b = \<mu> a + \<mu> (b - a)"
-    using measure_additive[of a "b - a"] Diff[of b a] assms by auto
-  moreover have "\<mu> a + 0 \<le> \<mu> a + \<mu> (b - a)" using assms by (intro add_mono) auto
-  ultimately show "\<mu> a \<le> \<mu> b" by auto
-qed
-
-lemma (in measure_space) measure_top:
-  "A \<in> sets M \<Longrightarrow> \<mu> A \<le> \<mu> (space M)"
-  using sets_into_space[of A] by (auto intro!: measure_mono)
-
-lemma (in measure_space) measure_compl:
-  assumes s: "s \<in> sets M" and fin: "\<mu> s \<noteq> \<infinity>"
-  shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
-proof -
-  have s_less_space: "\<mu> s \<le> \<mu> (space M)"
-    using s by (auto intro!: measure_mono sets_into_space)
-  from s have "0 \<le> \<mu> s" by auto
-  have "\<mu> (space M) = \<mu> (s \<union> (space M - s))" using s
-    by (metis Un_Diff_cancel Un_absorb1 s sets_into_space)
-  also have "... = \<mu> s + \<mu> (space M - s)"
-    by (rule additiveD [OF additive]) (auto simp add: s)
-  finally have "\<mu> (space M) = \<mu> s + \<mu> (space M - s)" .
-  then show ?thesis
-    using fin `0 \<le> \<mu> s`
-    unfolding ereal_eq_minus_iff by (auto simp: ac_simps)
-qed
-
-lemma (in measure_space) measure_Diff:
-  assumes finite: "\<mu> B \<noteq> \<infinity>"
-  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
-  shows "\<mu> (A - B) = \<mu> A - \<mu> B"
-proof -
-  have "0 \<le> \<mu> B" using assms by auto
-  have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
-  then have "\<mu> A = \<mu> ((A - B) \<union> B)" by simp
-  also have "\<dots> = \<mu> (A - B) + \<mu> B"
-    using measurable by (subst measure_additive[symmetric]) auto
-  finally show "\<mu> (A - B) = \<mu> A - \<mu> B"
-    unfolding ereal_eq_minus_iff
-    using finite `0 \<le> \<mu> B` by auto
-qed
-
-lemma (in measure_space) measure_countable_increasing:
-  assumes A: "range A \<subseteq> sets M"
-      and A0: "A 0 = {}"
-      and ASuc: "\<And>n. A n \<subseteq> A (Suc n)"
-  shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
-proof -
-  { fix n
-    have "\<mu> (A n) = (\<Sum>i<n. \<mu> (A (Suc i) - A i))"
-      proof (induct n)
-        case 0 thus ?case by (auto simp add: A0)
-      next
-        case (Suc m)
-        have "A (Suc m) = A m \<union> (A (Suc m) - A m)"
-          by (metis ASuc Un_Diff_cancel Un_absorb1)
-        hence "\<mu> (A (Suc m)) =
-               \<mu> (A m) + \<mu> (A (Suc m) - A m)"
-          by (subst measure_additive)
-             (auto simp add: measure_additive range_subsetD [OF A])
-        with Suc show ?case
-          by simp
-      qed }
-  note Meq = this
-  have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)"
-    proof (rule UN_finite2_eq [where k=1], simp)
-      fix i
-      show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)"
-        proof (induct i)
-          case 0 thus ?case by (simp add: A0)
-        next
-          case (Suc i)
-          thus ?case
-            by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc])
-        qed
-    qed
-  have A1: "\<And>i. A (Suc i) - A i \<in> sets M"
-    by (metis A Diff range_subsetD)
-  have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
-    by (blast intro: range_subsetD [OF A])
-  have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = (\<Sum>i. \<mu> (A (Suc i) - A i))"
-    using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric])
-  also have "\<dots> = \<mu> (\<Union>i. A (Suc i) - A i)"
-    by (rule measure_countably_additive)
-       (auto simp add: disjoint_family_Suc ASuc A1 A2)
-  also have "... =  \<mu> (\<Union>i. A i)"
-    by (simp add: Aeq)
-  finally have "(SUP n. \<Sum>i<n. \<mu> (A (Suc i) - A i)) = \<mu> (\<Union>i. A i)" .
-  then show ?thesis by (auto simp add: Meq)
-qed
-
-lemma (in measure_space) continuity_from_below:
-  assumes A: "range A \<subseteq> sets M" and "incseq A"
-  shows "(SUP n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
-proof -
-  have *: "(SUP n. \<mu> (nat_case {} A (Suc n))) = (SUP n. \<mu> (nat_case {} A n))"
-    using A by (auto intro!: SUPR_eq exI split: nat.split)
-  have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
-    by (auto simp add: split: nat.splits)
-  have meq: "\<And>n. \<mu> (A n) = (\<mu> \<circ> nat_case {} A) (Suc n)"
-    by simp
-  have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. nat_case {} A i)"
-    using range_subsetD[OF A] incseq_SucD[OF `incseq A`]
-    by (force split: nat.splits intro!: measure_countable_increasing)
-  also have "\<mu> (\<Union>i. nat_case {} A i) = \<mu> (\<Union>i. A i)"
-    by (simp add: ueq)
-  finally have "(SUP n. \<mu> (nat_case {} A n)) = \<mu> (\<Union>i. A i)" .
-  thus ?thesis unfolding meq * comp_def .
-qed
-
-lemma (in measure_space) measure_incseq:
-  assumes "range B \<subseteq> sets M" "incseq B"
-  shows "incseq (\<lambda>i. \<mu> (B i))"
-  using assms by (auto simp: incseq_def intro!: measure_mono)
-
-lemma (in measure_space) continuity_from_below_Lim:
-  assumes A: "range A \<subseteq> sets M" "incseq A"
-  shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Union>i. A i)"
-  using LIMSEQ_ereal_SUPR[OF measure_incseq, OF A]
-    continuity_from_below[OF A] by simp
-
-lemma (in measure_space) measure_decseq:
-  assumes "range B \<subseteq> sets M" "decseq B"
-  shows "decseq (\<lambda>i. \<mu> (B i))"
-  using assms by (auto simp: decseq_def intro!: measure_mono)
-
-lemma (in measure_space) continuity_from_above:
-  assumes A: "range A \<subseteq> sets M" and "decseq A"
-  and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
-  shows "(INF n. \<mu> (A n)) = \<mu> (\<Inter>i. A i)"
-proof -
-  have le_MI: "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
-    using A by (auto intro!: measure_mono)
-  hence *: "\<mu> (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto
-
-  have A0: "0 \<le> \<mu> (A 0)" using A by auto
-
-  have "\<mu> (A 0) - (INF n. \<mu> (A n)) = \<mu> (A 0) + (SUP n. - \<mu> (A n))"
-    by (simp add: ereal_SUPR_uminus minus_ereal_def)
-  also have "\<dots> = (SUP n. \<mu> (A 0) - \<mu> (A n))"
-    unfolding minus_ereal_def using A0 assms
-    by (subst SUPR_ereal_add) (auto simp add: measure_decseq)
-  also have "\<dots> = (SUP n. \<mu> (A 0 - A n))"
-    using A finite `decseq A`[unfolded decseq_def] by (subst measure_Diff) auto
-  also have "\<dots> = \<mu> (\<Union>i. A 0 - A i)"
-  proof (rule continuity_from_below)
-    show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
-      using A by auto
-    show "incseq (\<lambda>n. A 0 - A n)"
-      using `decseq A` by (auto simp add: incseq_def decseq_def)
-  qed
-  also have "\<dots> = \<mu> (A 0) - \<mu> (\<Inter>i. A i)"
-    using A finite * by (simp, subst measure_Diff) auto
-  finally show ?thesis
-    unfolding ereal_minus_eq_minus_iff using finite A0 by auto
-qed
-
-lemma (in measure_space) measure_insert:
-  assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
-  shows "\<mu> (insert x A) = \<mu> {x} + \<mu> A"
-proof -
-  have "{x} \<inter> A = {}" using `x \<notin> A` by auto
-  from measure_additive[OF sets this] show ?thesis by simp
-qed
-
-lemma (in measure_space) measure_setsum:
-  assumes "finite S" and "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
-  assumes disj: "disjoint_family_on A S"
-  shows "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>i\<in>S. A i)"
-using assms proof induct
-  case (insert i S)
-  then have "(\<Sum>i\<in>S. \<mu> (A i)) = \<mu> (\<Union>a\<in>S. A a)"
-    by (auto intro: disjoint_family_on_mono)
-  moreover have "A i \<inter> (\<Union>a\<in>S. A a) = {}"
-    using `disjoint_family_on A (insert i S)` `i \<notin> S`
-    by (auto simp: disjoint_family_on_def)
-  ultimately show ?case using insert
-    by (auto simp: measure_additive finite_UN)
-qed simp
-
-lemma (in measure_space) measure_finite_singleton:
-  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
-  shows "\<mu> S = (\<Sum>x\<in>S. \<mu> {x})"
-  using measure_setsum[of S "\<lambda>x. {x}", OF assms]
-  by (auto simp: disjoint_family_on_def)
-
-lemma finite_additivity_sufficient:
-  assumes "sigma_algebra M"
-  assumes fin: "finite (space M)" and pos: "positive M (measure M)" and add: "additive M (measure M)"
-  shows "measure_space M"
-proof -
-  interpret sigma_algebra M by fact
-  show ?thesis
-  proof
-    show [simp]: "positive M (measure M)" using pos by (simp add: positive_def)
-    show "countably_additive M (measure M)"
-    proof (auto simp add: countably_additive_def)
-      fix A :: "nat \<Rightarrow> 'a set"
-      assume A: "range A \<subseteq> sets M"
-         and disj: "disjoint_family A"
-         and UnA: "(\<Union>i. A i) \<in> sets M"
-      def I \<equiv> "{i. A i \<noteq> {}}"
-      have "Union (A ` I) \<subseteq> space M" using A
-        by auto (metis range_subsetD subsetD sets_into_space)
-      hence "finite (A ` I)"
-        by (metis finite_UnionD finite_subset fin)
-      moreover have "inj_on A I" using disj
-        by (auto simp add: I_def disjoint_family_on_def inj_on_def)
-      ultimately have finI: "finite I"
-        by (metis finite_imageD)
-      hence "\<exists>N. \<forall>m\<ge>N. A m = {}"
-        proof (cases "I = {}")
-          case True thus ?thesis by (simp add: I_def)
-        next
-          case False
-          hence "\<forall>i\<in>I. i < Suc(Max I)"
-            by (simp add: Max_less_iff [symmetric] finI)
-          hence "\<forall>m \<ge> Suc(Max I). A m = {}"
-            by (simp add: I_def) (metis less_le_not_le)
-          thus ?thesis
-            by blast
-        qed
-      then obtain N where N: "\<forall>m\<ge>N. A m = {}" by blast
-      then have "\<forall>m\<ge>N. measure M (A m) = 0" using pos[unfolded positive_def] by simp
-      then have "(\<Sum>n. measure M (A n)) = (\<Sum>m<N. measure M (A m))"
-        by (simp add: suminf_finite)
-      also have "... = measure M (\<Union>i<N. A i)"
-        proof (induct N)
-          case 0 thus ?case using pos[unfolded positive_def] by simp
-        next
-          case (Suc n)
-          have "measure M (A n \<union> (\<Union> x<n. A x)) = measure M (A n) + measure M (\<Union> i<n. A i)"
-            proof (rule Caratheodory.additiveD [OF add])
-              show "A n \<inter> (\<Union> x<n. A x) = {}" using disj
-                by (auto simp add: disjoint_family_on_def nat_less_le) blast
-              show "A n \<in> sets M" using A
-                by force
-              show "(\<Union>i<n. A i) \<in> sets M"
-                proof (induct n)
-                  case 0 thus ?case by simp
-                next
-                  case (Suc n) thus ?case using A
-                    by (simp add: lessThan_Suc Un range_subsetD)
-                qed
-            qed
-          thus ?case using Suc
-            by (simp add: lessThan_Suc)
-        qed
-      also have "... = measure M (\<Union>i. A i)"
-        proof -
-          have "(\<Union> i<N. A i) = (\<Union>i. A i)" using N
-            by auto (metis Int_absorb N disjoint_iff_not_equal lessThan_iff not_leE)
-          thus ?thesis by simp
-        qed
-      finally show "(\<Sum>n. measure M (A n)) = measure M (\<Union>i. A i)" .
-    qed
-  qed
-qed
-
-lemma (in measure_space) measure_setsum_split:
-  assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
-  assumes "(\<Union>i\<in>S. B i) = space M"
-  assumes "disjoint_family_on B S"
-  shows "\<mu> A = (\<Sum>i\<in>S. \<mu> (A \<inter> (B i)))"
-proof -
-  have *: "\<mu> A = \<mu> (\<Union>i\<in>S. A \<inter> B i)"
-    using assms by auto
-  show ?thesis unfolding *
-  proof (rule measure_setsum[symmetric])
-    show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
-      using `disjoint_family_on B S`
-      unfolding disjoint_family_on_def by auto
-  qed (insert assms, auto)
-qed
-
-lemma (in measure_space) measure_subadditive:
-  assumes measurable: "A \<in> sets M" "B \<in> sets M"
-  shows "\<mu> (A \<union> B) \<le> \<mu> A + \<mu> B"
-proof -
-  from measure_additive[of A "B - A"]
-  have "\<mu> (A \<union> B) = \<mu> A + \<mu> (B - A)"
-    using assms by (simp add: Diff)
-  also have "\<dots> \<le> \<mu> A + \<mu> B"
-    using assms by (auto intro!: add_left_mono measure_mono)
-  finally show ?thesis .
-qed
-
-lemma (in measure_space) measure_subadditive_finite:
-  assumes "finite I" "\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets M"
-  shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))"
-using assms proof induct
-  case (insert i I)
-  then have "\<mu> (\<Union>i\<in>insert i I. A i) = \<mu> (A i \<union> (\<Union>i\<in>I. A i))"
-    by simp
-  also have "\<dots> \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)"
-    using insert by (intro measure_subadditive finite_UN) auto
-  also have "\<dots> \<le> \<mu> (A i) + (\<Sum>i\<in>I. \<mu> (A i))"
-    using insert by (intro add_mono) auto
-  also have "\<dots> = (\<Sum>i\<in>insert i I. \<mu> (A i))"
-    using insert by auto
-  finally show ?case .
-qed simp
-
-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) positive_measure[OF assms(4)] 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))"
-using assms proof induct
-  case (insert i I)
-  then have "(\<Union>i\<in>I. A i) \<in> sets M" by auto
-  then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)"
-    using insert by (simp add: measure_subadditive)
-  also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))"
-    using insert by (auto intro!: add_left_mono)
-  finally show ?case .
-qed simp
-
-lemma (in measure_space) measure_countably_subadditive:
-  assumes "range f \<subseteq> sets M"
-  shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>i. \<mu> (f i))"
-proof -
-  have "\<mu> (\<Union>i. f i) = \<mu> (\<Union>i. disjointed f i)"
-    unfolding UN_disjointed_eq ..
-  also have "\<dots> = (\<Sum>i. \<mu> (disjointed f i))"
-    using range_disjointed_sets[OF assms] measure_countably_additive
-    by (simp add:  disjoint_family_disjointed comp_def)
-  also have "\<dots> \<le> (\<Sum>i. \<mu> (f i))"
-    using range_disjointed_sets[OF assms] assms
-    by (auto intro!: suminf_le_pos measure_mono disjointed_subset)
-  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"
-proof -
-  have "0 \<le> \<mu> (\<Union> i. N i)" using assms by auto
-  moreover have "\<mu> (\<Union> i. N i) \<le> 0"
-    using measure_countably_subadditive[OF assms(2)] assms(1) by simp
-  ultimately show ?thesis by simp
-qed
-
-lemma (in measure_space) measure_inter_full_set:
-  assumes "S \<in> sets M" "T \<in> sets M" and fin: "\<mu> (T - S) \<noteq> \<infinity>"
-  assumes T: "\<mu> T = \<mu> (space M)"
-  shows "\<mu> (S \<inter> T) = \<mu> S"
-proof (rule antisym)
-  show " \<mu> (S \<inter> T) \<le> \<mu> S"
-    using assms by (auto intro!: measure_mono)
-
-  have pos: "0 \<le> \<mu> (T - S)" using assms by auto
-  show "\<mu> S \<le> \<mu> (S \<inter> T)"
-  proof (rule ccontr)
-    assume contr: "\<not> ?thesis"
-    have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))"
-      unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"])
-    also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)"
-      using assms by (auto intro!: measure_subadditive)
-    also have "\<dots> < \<mu> (T - S) + \<mu> S"
-      using fin contr pos by (intro ereal_less_add) auto
-    also have "\<dots> = \<mu> (T \<union> S)"
-      using assms by (subst measure_additive) auto
-    also have "\<dots> \<le> \<mu> (space M)"
-      using assms sets_into_space by (auto intro!: measure_mono)
-    finally show False ..
-  qed
-qed
-
-lemma measure_unique_Int_stable:
-  fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \<Rightarrow> 'a set"
-  assumes "Int_stable E"
-  and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E"
-  and M: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<mu>\<rparr>" (is "measure_space ?M")
-  and N: "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<nu>\<rparr>" (is "measure_space ?N")
-  and eq: "\<And>X. X \<in> sets E \<Longrightarrow> \<mu> X = \<nu> X"
-  and finite: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
-  assumes "X \<in> sets (sigma E)"
-  shows "\<mu> X = \<nu> X"
-proof -
-  let ?D = "\<lambda>F. {D. D \<in> sets (sigma E) \<and> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)}"
-  interpret M: measure_space ?M
-    where "space ?M = space E" and "sets ?M = sets (sigma E)" and "measure ?M = \<mu>" by (simp_all add: M)
-  interpret N: measure_space ?N
-    where "space ?N = space E" and "sets ?N = sets (sigma E)" and "measure ?N = \<nu>" by (simp_all add: N)
-  { fix F assume "F \<in> sets E" and "\<mu> F \<noteq> \<infinity>"
-    then have [intro]: "F \<in> sets (sigma E)" by auto
-    have "\<nu> F \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` `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 (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
-      then show "A \<subseteq> space E" using M.sets_into_space by auto
-    next
-      have "F \<inter> space E = F" using `F \<in> sets E` by auto
-      then show "\<mu> (F \<inter> space E) = \<nu> (F \<inter> space E)"
-        using `F \<in> sets E` eq by auto
-    next
-      fix A assume *: "A \<in> sets (sigma E) \<and> \<mu> (F \<inter> A) = \<nu> (F \<inter> A)"
-      then have **: "F \<inter> (space (sigma E) - A) = F - (F \<inter> A)"
-        and [intro]: "F \<inter> A \<in> sets (sigma E)"
-        using `F \<in> sets E` M.sets_into_space by auto
-      have "\<nu> (F \<inter> A) \<le> \<nu> F" by (auto intro!: N.measure_mono)
-      then have "\<nu> (F \<inter> A) \<noteq> \<infinity>" using `\<nu> F \<noteq> \<infinity>` by auto
-      have "\<mu> (F \<inter> A) \<le> \<mu> F" by (auto intro!: M.measure_mono)
-      then have "\<mu> (F \<inter> A) \<noteq> \<infinity>" using `\<mu> F \<noteq> \<infinity>` by auto
-      then have "\<mu> (F \<inter> (space (sigma E) - A)) = \<mu> F - \<mu> (F \<inter> A)" unfolding **
-        using `F \<inter> A \<in> sets (sigma E)` 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 (sigma E) - A))" unfolding **
-        using `F \<inter> A \<in> sets (sigma E)` `\<nu> (F \<inter> A) \<noteq> \<infinity>` by (auto intro!: N.measure_Diff[symmetric])
-      finally show "space E - A \<in> sets (sigma E) \<and> \<mu> (F \<inter> (space E - A)) = \<nu> (F \<inter> (space E - A))"
-        using * by auto
-    next
-      fix A :: "nat \<Rightarrow> 'a set"
-      assume "disjoint_family A" "range A \<subseteq> {X \<in> sets (sigma E). \<mu> (F \<inter> X) = \<nu> (F \<inter> X)}"
-      then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sets (sigma E)" "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 (sigma E)"
-        by (auto simp: disjoint_family_on_def subset_eq)
-      then show "(\<Union>x. A x) \<in> sets (sigma E) \<and> \<mu> (F \<inter> (\<Union>x. A x)) = \<nu> (F \<inter> (\<Union>x. A x))"
-        by (auto simp: M.measure_countably_additive[symmetric]
-                       N.measure_countably_additive[symmetric]
-            simp del: UN_simps)
-    qed
-    have *: "sets (sigma E) = sets \<lparr>space = space E, sets = ?D F\<rparr>"
-      using `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 (sigma E) \<Longrightarrow> \<mu> (F \<inter> D) = \<nu> (F \<inter> D)"
-      by (subst (asm) *) auto }
-  note * = this
-  let ?A = "\<lambda>i. A i \<inter> X"
-  have A': "range ?A \<subseteq> sets (sigma E)" "incseq ?A"
-    using A(1,2) `X \<in> sets (sigma E)` by (auto simp: incseq_def)
-  { fix i have "\<mu> (?A i) = \<nu> (?A i)"
-      using *[of "A i" X] `X \<in> sets (sigma E)` A finite by auto }
-  with M.continuity_from_below[OF A'] N.continuity_from_below[OF A']
-  show ?thesis using A(3) `X \<in> sets (sigma E)` by auto
-qed
-
-section "@{text \<mu>}-null sets"
-
-abbreviation (in measure_space) "null_sets \<equiv> {N\<in>sets M. \<mu> N = 0}"
-
-sublocale measure_space \<subseteq> nullsets!: ring_of_sets "\<lparr> space = space M, sets = null_sets \<rparr>"
-  where "space \<lparr> space = space M, sets = null_sets \<rparr> = space M"
-  and "sets \<lparr> space = space M, sets = null_sets \<rparr> = null_sets"
-proof -
-  { fix A B assume sets: "A \<in> sets M" "B \<in> sets M"
-    moreover then have "\<mu> (A \<union> B) \<le> \<mu> A + \<mu> B" "\<mu> (A - B) \<le> \<mu> A"
-      by (auto intro!: measure_subadditive measure_mono)
-    moreover assume "\<mu> B = 0" "\<mu> A = 0"
-    ultimately have "\<mu> (A - B) = 0" "\<mu> (A \<union> B) = 0"
-      by (auto intro!: antisym) }
-  note null = this
-  show "ring_of_sets \<lparr> space = space M, sets = null_sets \<rparr>"
-    by default (insert sets_into_space null, auto)
-qed simp_all
-
-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 SUP_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
-  then have "0 \<le> \<mu> (\<Union>i. N i)" by simp
-  moreover have "\<mu> (\<Union>i. N i) \<le> (\<Sum>n. \<mu> (N (Countable.from_nat n)))"
-    unfolding UN_from_nat[of N]
-    using assms by (intro measure_countably_subadditive) auto
-  ultimately 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
-
-section "Formalise almost everywhere"
-
-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)"
-
-syntax
-  "_almost_everywhere" :: "pttrn \<Rightarrow> ('a, 'b) measure_space_scheme \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
-
-translations
-  "AE x in M. P" == "CONST measure_space.almost_everywhere M (%x. P)"
-
-lemma (in measure_space) AE_cong_measure:
-  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "sets N = sets M" "space N = space M"
-  shows "(AE x in N. P x) \<longleftrightarrow> (AE x. P x)"
-proof -
-  interpret N: measure_space N
-    by (rule measure_space_cong) fact+
-  show ?thesis
-    unfolding N.almost_everywhere_def almost_everywhere_def
-    by (auto simp: assms)
-qed
-
-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
-  have "0 \<le> \<mu> ?P" using assms by simp
-  moreover have "\<mu> ?P \<le> \<mu> N"
-    using assms N(1,2) by (auto intro: measure_mono)
-  ultimately have "\<mu> ?P = 0" unfolding `\<mu> N = 0` by auto
-  then show "?P \<in> null_sets" using assms by simp
-next
-  assume "?P \<in> null_sets" with assms show "AE x. P x" by (auto intro: AE_I')
-qed
-
-lemma (in measure_space) AE_iff_measurable:
-  "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x. P x) \<longleftrightarrow> \<mu> N = 0"
-  using AE_iff_null_set[of P] by simp
-
-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_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 -
-  have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}"
-    by auto
-  with AE_iff_null_set[of P] assms show ?thesis by auto
-qed
-
-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[elim!]:
-  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)
-    have "0 \<le> \<mu> (A \<union> B)" using A B by auto
-    moreover have "\<mu> (A \<union> B) \<le> 0"
-      using measure_subadditive[of A B] A B by auto
-    ultimately show "A \<union> B \<in> sets M" "\<mu> (A \<union> B) = 0" using A B by auto
-    show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
-      using P imp by auto
-  qed
-qed
-
-lemma (in measure_space)
-  shows AE_iffI: "AE x. P x \<Longrightarrow> AE x. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x. Q x"
-    and AE_disjI1: "AE x. P x \<Longrightarrow> AE x. P x \<or> Q x"
-    and AE_disjI2: "AE x. Q x \<Longrightarrow> AE x. P x \<or> Q x"
-    and AE_conjI: "AE x. P x \<Longrightarrow> AE x. Q x \<Longrightarrow> AE x. P x \<and> Q x"
-    and AE_conj_iff[simp]: "(AE x. P x \<and> Q x) \<longleftrightarrow> (AE x. P x) \<and> (AE x. Q x)"
-  by auto
-
-lemma (in measure_space) AE_measure:
-  assumes AE: "AE x. P x" and sets: "{x\<in>space M. P x} \<in> sets M"
-  shows "\<mu> {x\<in>space M. P x} = \<mu> (space M)"
-proof -
-  from AE_E[OF AE] guess N . note N = this
-  with sets have "\<mu> (space M) \<le> \<mu> ({x\<in>space M. P x} \<union> N)"
-    by (intro measure_mono) auto
-  also have "\<dots> \<le> \<mu> {x\<in>space M. P x} + \<mu> N"
-    using sets N by (intro measure_subadditive) auto
-  also have "\<dots> = \<mu> {x\<in>space M. P x}" using N by simp
-  finally show "\<mu> {x\<in>space M. P x} = \<mu> (space M)"
-    using measure_top[OF sets] by auto
-qed
-
-lemma (in measure_space) AE_space: "AE x. x \<in> space M"
-  by (rule AE_I[where N="{}"]) auto
-
-lemma (in measure_space) AE_I2[simp, intro]:
-  "(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x. P x"
-  using AE_space by auto
-
-lemma (in measure_space) AE_Ball_mp:
-  "\<forall>x\<in>space M. P x \<Longrightarrow> AE x. P x \<longrightarrow> Q x \<Longrightarrow> AE x. Q x"
-  by auto
-
-lemma (in measure_space) AE_cong[cong]:
-  "(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x. P x) \<longleftrightarrow> (AE x. Q x)"
-  by auto
-
-lemma (in measure_space) AE_all_countable:
-  "(AE x. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x. 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
-qed auto
-
-lemma (in measure_space) AE_finite_all:
-  assumes f: "finite S" shows "(AE x. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x. P i x)"
-  using f by induct auto
-
-lemma (in measure_space) restricted_measure_space:
-  assumes "S \<in> sets M"
-  shows "measure_space (restricted_space S)"
-    (is "measure_space ?r")
-  unfolding measure_space_def measure_space_axioms_def
-proof safe
-  show "sigma_algebra ?r" using restricted_sigma_algebra[OF assms] .
-  show "positive ?r (measure ?r)" using `S \<in> sets M` by (auto simp: positive_def)
-
-  show "countably_additive ?r (measure ?r)"
-    unfolding countably_additive_def
-  proof safe
-    fix A :: "nat \<Rightarrow> 'a set"
-    assume *: "range A \<subseteq> sets ?r" and **: "disjoint_family A"
-    from restriction_in_sets[OF assms *[simplified]] **
-    show "(\<Sum>n. measure ?r (A n)) = measure ?r (\<Union>i. A i)"
-      using measure_countably_additive by simp
-  qed
-qed
-
-lemma (in measure_space) AE_restricted:
-  assumes "A \<in> sets M"
-  shows "(AE x in restricted_space A. P x) \<longleftrightarrow> (AE x. x \<in> A \<longrightarrow> P x)"
-proof -
-  interpret R: measure_space "restricted_space A"
-    by (rule restricted_measure_space[OF `A \<in> sets M`])
-  show ?thesis
-  proof
-    assume "AE x in restricted_space A. P x"
-    from this[THEN R.AE_E] guess N' .
-    then obtain N where "{x \<in> A. \<not> P x} \<subseteq> A \<inter> N" "\<mu> (A \<inter> N) = 0" "N \<in> sets M"
-      by auto
-    moreover then have "{x \<in> space M. \<not> (x \<in> A \<longrightarrow> P x)} \<subseteq> A \<inter> N"
-      using `A \<in> sets M` sets_into_space by auto
-    ultimately show "AE x. x \<in> A \<longrightarrow> P x"
-      using `A \<in> sets M` by (auto intro!: AE_I[where N="A \<inter> N"])
-  next
-    assume "AE x. x \<in> A \<longrightarrow> P x"
-    from this[THEN AE_E] guess N .
-    then show "AE x in restricted_space A. P x"
-      using null_set_Int1[OF _ `A \<in> sets M`] `A \<in> sets M`[THEN sets_into_space]
-      by (auto intro!: R.AE_I[where N="A \<inter> N"] simp: subset_eq)
-  qed
-qed
-
-lemma (in measure_space) measure_space_subalgebra:
-  assumes "sigma_algebra N" and "sets N \<subseteq> sets M" "space N = space M"
-  and measure[simp]: "\<And>X. X \<in> sets N \<Longrightarrow> measure N X = measure M X"
-  shows "measure_space N"
-proof -
-  interpret N: sigma_algebra N by fact
-  show ?thesis
-  proof
-    from `sets N \<subseteq> sets M` have "\<And>A. range A \<subseteq> sets N \<Longrightarrow> range A \<subseteq> sets M" by blast
-    then show "countably_additive N (measure N)"
-      by (auto intro!: measure_countably_additive simp: countably_additive_def subset_eq)
-    show "positive N (measure_space.measure N)"
-      using assms(2) by (auto simp add: positive_def)
-  qed
-qed
-
-lemma (in measure_space) AE_subalgebra:
-  assumes ae: "AE x in N. P x"
-  and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
-  and sa: "sigma_algebra N"
-  shows "AE x. P x"
-proof -
-  interpret N: measure_space N using measure_space_subalgebra[OF sa N] .
-  from ae[THEN N.AE_E] guess N .
-  with N show ?thesis unfolding almost_everywhere_def by auto
-qed
-
-section "@{text \<sigma>}-finite Measures"
-
-locale sigma_finite_measure = measure_space +
-  assumes 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> \<infinity>)"
-
-lemma (in sigma_finite_measure) restricted_sigma_finite_measure:
-  assumes "S \<in> sets M"
-  shows "sigma_finite_measure (restricted_space S)"
-    (is "sigma_finite_measure ?r")
-  unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def
-proof safe
-  show "measure_space ?r" using restricted_measure_space[OF assms] .
-next
-  obtain A :: "nat \<Rightarrow> 'a set" where
-      "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
-    using sigma_finite by auto
-  show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets ?r \<and> (\<Union>i. A i) = space ?r \<and> (\<forall>i. measure ?r (A i) \<noteq> \<infinity>)"
-  proof (safe intro!: exI[of _ "\<lambda>i. A i \<inter> S"] del: notI)
-    fix i
-    show "A i \<inter> S \<in> sets ?r"
-      using `range A \<subseteq> sets M` `S \<in> sets M` by auto
-  next
-    fix x i assume "x \<in> S" thus "x \<in> space ?r" by simp
-  next
-    fix x assume "x \<in> space ?r" thus "x \<in> (\<Union>i. A i \<inter> S)"
-      using `(\<Union>i. A i) = space M` `S \<in> sets M` by auto
-  next
-    fix i
-    have "\<mu> (A i \<inter> S) \<le> \<mu> (A i)"
-      using `range A \<subseteq> sets M` `S \<in> sets M` by (auto intro!: measure_mono)
-    then show "measure ?r (A i \<inter> S) \<noteq> \<infinity>" using `\<mu> (A i) \<noteq> \<infinity>` by auto
-  qed
-qed
-
-lemma (in sigma_finite_measure) sigma_finite_measure_cong:
-  assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> measure M' A = \<mu> A" "sets M' = sets M" "space M' = space M"
-  shows "sigma_finite_measure M'"
-proof -
-  interpret M': measure_space M' by (intro measure_space_cong 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. measure M' (A i) \<noteq> \<infinity>" using cong by auto
-  show ?thesis
-    apply default
-    using A fin cong 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> \<infinity>) \<and> disjoint_family A"
-proof -
-  obtain A :: "nat \<Rightarrow> 'a set" where
-    range: "range A \<subseteq> sets M" and
-    space: "(\<Union>i. A i) = space M" and
-    measure: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
-    using sigma_finite by auto
-  note range' = range_disjointed_sets[OF range] range
-  { fix i
-    have "\<mu> (disjointed A i) \<le> \<mu> (A i)"
-      using range' disjointed_subset[of A i] by (auto intro!: measure_mono)
-    then have "\<mu> (disjointed A i) \<noteq> \<infinity>"
-      using measure[of i] by auto }
-  with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
-  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> incseq F \<and> (\<Union>i. F i) = space M \<and> (\<forall>i. \<mu> (F i) \<noteq> \<infinity>)"
-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> \<infinity>"
-    using sigma_finite by auto
-  then show ?thesis
-  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 fastforce
-  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> < \<infinity>"
-      using F by (auto simp: setsum_Pinfty)
-    finally show "\<mu> (\<Union> i\<le>n. F i) \<noteq> \<infinity>" by simp
-  qed (force simp: incseq_def)+
-qed
-
-section {* Measure preserving *}
-
-definition "measure_preserving A B =
-    {f \<in> measurable A B. (\<forall>y \<in> sets B. measure B y = measure A (f -` y \<inter> space A))}"
-
-lemma measure_preservingI[intro?]:
-  assumes "f \<in> measurable A B"
-    and "\<And>y. y \<in> sets B \<Longrightarrow> measure A (f -` y \<inter> space A) = measure B y"
-  shows "f \<in> measure_preserving A B"
-  unfolding measure_preserving_def using assms by auto
-
-lemma (in measure_space) measure_space_vimage:
-  fixes M' :: "('c, 'd) measure_space_scheme"
-  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
-  shows "measure_space M'"
-proof -
-  interpret M': sigma_algebra M' by fact
-  show ?thesis
-  proof
-    show "positive M' (measure M')" using T
-      by (auto simp: measure_preserving_def positive_def measurable_sets)
-
-    show "countably_additive M' (measure M')"
-    proof (intro countably_additiveI)
-      fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets M'" "disjoint_family A"
-      then have A: "\<And>i. A i \<in> sets M'" "(\<Union>i. A i) \<in> sets M'" by auto
-      then have *: "range (\<lambda>i. T -` (A i) \<inter> space M) \<subseteq> sets M"
-        using T by (auto simp: measurable_def measure_preserving_def)
-      moreover have "(\<Union>i. T -`  A i \<inter> space M) \<in> sets M"
-        using * by blast
-      moreover have **: "disjoint_family (\<lambda>i. T -` A i \<inter> space M)"
-        using `disjoint_family A` by (auto simp: disjoint_family_on_def)
-      ultimately show "(\<Sum>i. measure M' (A i)) = measure M' (\<Union>i. A i)"
-        using measure_countably_additive[OF _ **] A T
-        by (auto simp: comp_def vimage_UN measure_preserving_def)
-    qed
-  qed
-qed
-
-lemma (in measure_space) almost_everywhere_vimage:
-  assumes T: "sigma_algebra M'" "T \<in> measure_preserving M M'"
-    and AE: "measure_space.almost_everywhere M' P"
-  shows "AE x. P (T x)"
-proof -
-  interpret M': measure_space M' using T by (rule measure_space_vimage)
-  from AE[THEN M'.AE_E] guess N .
-  then show ?thesis
-    unfolding almost_everywhere_def M'.almost_everywhere_def
-    using T(2) unfolding measurable_def measure_preserving_def
-    by (intro bexI[of _ "T -` N \<inter> space M"]) auto
-qed
-
-lemma measure_unique_Int_stable_vimage:
-  fixes A :: "nat \<Rightarrow> 'a set"
-  assumes E: "Int_stable E"
-  and A: "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure M (A i) \<noteq> \<infinity>"
-  and N: "measure_space N" "T \<in> measurable N M"
-  and M: "measure_space M" "sets (sigma E) = sets M" "space E = space M"
-  and eq: "\<And>X. X \<in> sets E \<Longrightarrow> measure M X = measure N (T -` X \<inter> space N)"
-  assumes X: "X \<in> sets (sigma E)"
-  shows "measure M X = measure N (T -` X \<inter> space N)"
-proof (rule measure_unique_Int_stable[OF E A(1,2,3) _ _ eq _ X])
-  interpret M: measure_space M by fact
-  interpret N: measure_space N by fact
-  let ?T = "\<lambda>X. T -` X \<inter> space N"
-  show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = measure M\<rparr>"
-    by (rule M.measure_space_cong) (auto simp: M)
-  show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure = \<lambda>X. measure N (?T X)\<rparr>" (is "measure_space ?E")
-  proof (rule N.measure_space_vimage)
-    show "sigma_algebra ?E"
-      by (rule M.sigma_algebra_cong) (auto simp: M)
-    show "T \<in> measure_preserving N ?E"
-      using `T \<in> measurable N M` by (auto simp: M measurable_def measure_preserving_def)
-  qed
-  show "\<And>i. M.\<mu> (A i) \<noteq> \<infinity>" by fact
-qed
-
-lemma (in measure_space) measure_preserving_Int_stable:
-  fixes A :: "nat \<Rightarrow> 'a set"
-  assumes E: "Int_stable E" "range A \<subseteq> sets E" "incseq A" "(\<Union>i. A i) = space E" "\<And>i. measure E (A i) \<noteq> \<infinity>"
-  and N: "measure_space (sigma E)"
-  and T: "T \<in> measure_preserving M E"
-  shows "T \<in> measure_preserving M (sigma E)"
-proof
-  interpret E: measure_space "sigma E" by fact
-  show "T \<in> measurable M (sigma E)"
-    using T E.sets_into_space
-    by (intro measurable_sigma) (auto simp: measure_preserving_def measurable_def)
-  fix X assume "X \<in> sets (sigma E)"
-  show "\<mu> (T -` X \<inter> space M) = E.\<mu> X"
-  proof (rule measure_unique_Int_stable_vimage[symmetric])
-    show "sets (sigma E) = sets (sigma E)" "space E = space (sigma E)"
-      "\<And>i. E.\<mu> (A i) \<noteq> \<infinity>" using E by auto
-    show "measure_space M" by default
-  next
-    fix X assume "X \<in> sets E" then show "E.\<mu> X = \<mu> (T -` X \<inter> space M)"
-      using T unfolding measure_preserving_def by auto
-  qed fact+
-qed
-
-section "Real measure values"
-
-lemma (in measure_space) real_measure_Union:
-  assumes finite: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>"
-  and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
-  shows "real (\<mu> (A \<union> B)) = real (\<mu> A) + real (\<mu> B)"
-  unfolding measure_additive[symmetric, OF measurable]
-  using measurable(1,2)[THEN positive_measure]
-  using finite by (cases rule: ereal2_cases[of "\<mu> A" "\<mu> B"]) auto
-
-lemma (in measure_space) real_measure_finite_Union:
-  assumes measurable:
-    "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M" "disjoint_family_on A S"
-  assumes finite: "\<And>i. i \<in> S \<Longrightarrow> \<mu> (A i) \<noteq> \<infinity>"
-  shows "real (\<mu> (\<Union>i\<in>S. A i)) = (\<Sum>i\<in>S. real (\<mu> (A i)))"
-  using finite measurable(2)[THEN positive_measure]
-  by (force intro!: setsum_real_of_ereal[symmetric]
-            simp: measure_setsum[OF measurable, symmetric])
-
-lemma (in measure_space) real_measure_Diff:
-  assumes finite: "\<mu> A \<noteq> \<infinity>"
-  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
-  shows "real (\<mu> (A - B)) = real (\<mu> A) - real (\<mu> B)"
-proof -
-  have "\<mu> (A - B) \<le> \<mu> A" "\<mu> B \<le> \<mu> A"
-    using measurable by (auto intro!: measure_mono)
-  hence "real (\<mu> ((A - B) \<union> B)) = real (\<mu> (A - B)) + real (\<mu> B)"
-    using measurable finite by (rule_tac real_measure_Union) auto
-  thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2)
-qed
-
-lemma (in measure_space) real_measure_UNION:
-  assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
-  assumes finite: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>"
-  shows "(\<lambda>i. real (\<mu> (A i))) sums (real (\<mu> (\<Union>i. A i)))"
-proof -
-  have "\<And>i. 0 \<le> \<mu> (A i)" using measurable by auto
-  with summable_sums[OF summable_ereal_pos, of "\<lambda>i. \<mu> (A i)"]
-     measure_countably_additive[OF measurable]
-  have "(\<lambda>i. \<mu> (A i)) sums (\<mu> (\<Union>i. A i))" by simp
-  moreover
-  { fix i
-    have "\<mu> (A i) \<le> \<mu> (\<Union>i. A i)"
-      using measurable by (auto intro!: measure_mono)
-    moreover have "0 \<le> \<mu> (A i)" using measurable by auto
-    ultimately have "\<mu> (A i) = ereal (real (\<mu> (A i)))"
-      using finite by (cases "\<mu> (A i)") auto }
-  moreover
-  have "0 \<le> \<mu> (\<Union>i. A i)" using measurable by auto
-  then have "\<mu> (\<Union>i. A i) = ereal (real (\<mu> (\<Union>i. A i)))"
-    using finite by (cases "\<mu> (\<Union>i. A i)") auto
-  ultimately show ?thesis
-    unfolding sums_ereal[symmetric] by simp
-qed
-
-lemma (in measure_space) real_measure_subadditive:
-  assumes measurable: "A \<in> sets M" "B \<in> sets M"
-  and fin: "\<mu> A \<noteq> \<infinity>" "\<mu> B \<noteq> \<infinity>"
-  shows "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
-proof -
-  have "0 \<le> \<mu> (A \<union> B)" using measurable by auto
-  then show "real (\<mu> (A \<union> B)) \<le> real (\<mu> A) + real (\<mu> B)"
-    using measure_subadditive[OF measurable] fin
-    by (cases rule: ereal3_cases[of "\<mu> (A \<union> B)" "\<mu> A" "\<mu> B"]) auto
-qed
-
-lemma (in measure_space) real_measure_setsum_singleton:
-  assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
-  and fin: "\<And>x. x \<in> S \<Longrightarrow> \<mu> {x} \<noteq> \<infinity>"
-  shows "real (\<mu> S) = (\<Sum>x\<in>S. real (\<mu> {x}))"
-  using measure_finite_singleton[OF S] fin
-  using positive_measure[OF S(2)]
-  by (force intro!: setsum_real_of_ereal[symmetric])
-
-lemma (in measure_space) real_continuity_from_below:
-  assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "\<mu> (\<Union>i. A i) \<noteq> \<infinity>"
-  shows "(\<lambda>i. real (\<mu> (A i))) ----> real (\<mu> (\<Union>i. A i))"
-proof -
-  have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
-  then have "ereal (real (\<mu> (\<Union>i. A i))) = \<mu> (\<Union>i. A i)"
-    using fin by (auto intro: ereal_real')
-  then show ?thesis
-    using continuity_from_below_Lim[OF A]
-    by (intro lim_real_of_ereal) simp
-qed
-
-lemma (in measure_space) continuity_from_above_Lim:
-  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
-  shows "(\<lambda>i. (\<mu> (A i))) ----> \<mu> (\<Inter>i. A i)"
-  using LIMSEQ_ereal_INFI[OF measure_decseq, OF A]
-  using continuity_from_above[OF A fin] by simp
-
-lemma (in measure_space) real_continuity_from_above:
-  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. \<mu> (A i) \<noteq> \<infinity>"
-  shows "(\<lambda>n. real (\<mu> (A n))) ----> real (\<mu> (\<Inter>i. A i))"
-proof -
-  have "0 \<le> \<mu> (\<Inter>i. A i)" using A by auto
-  moreover
-  have "\<mu> (\<Inter>i. A i) \<le> \<mu> (A 0)"
-    using A by (auto intro!: measure_mono)
-  ultimately have "ereal (real (\<mu> (\<Inter>i. A i))) = \<mu> (\<Inter>i. A i)"
-    using fin by (auto intro: ereal_real')
-  then show ?thesis
-    using continuity_from_above_Lim[OF A fin]
-    by (intro lim_real_of_ereal) simp
-qed
-
-lemma (in measure_space) real_measure_countably_subadditive:
-  assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. \<mu> (A i)) \<noteq> \<infinity>"
-  shows "real (\<mu> (\<Union>i. A i)) \<le> (\<Sum>i. real (\<mu> (A i)))"
-proof -
-  { fix i
-    have "0 \<le> \<mu> (A i)" using A by auto
-    moreover have "\<mu> (A i) \<noteq> \<infinity>" using A by (intro suminf_PInfty[OF _ fin]) auto
-    ultimately have "\<bar>\<mu> (A i)\<bar> \<noteq> \<infinity>" by auto }
-  moreover have "0 \<le> \<mu> (\<Union>i. A i)" using A by auto
-  ultimately have "ereal (real (\<mu> (\<Union>i. A i))) \<le> (\<Sum>i. ereal (real (\<mu> (A i))))"
-    using measure_countably_subadditive[OF A] by (auto simp: ereal_real)
-  also have "\<dots> = ereal (\<Sum>i. real (\<mu> (A i)))"
-    using A
-    by (auto intro!: sums_unique[symmetric] sums_ereal[THEN iffD2] summable_sums summable_real_of_ereal fin)
-  finally show ?thesis by simp
-qed
-
-locale finite_measure = sigma_finite_measure +
-  assumes finite_measure_of_space: "\<mu> (space M) \<noteq> \<infinity>"
-
-lemma finite_measureI[Pure.intro!]:
-  assumes "measure_space M"
-  assumes *: "measure M (space M) \<noteq> \<infinity>"
-  shows "finite_measure M"
-proof -
-  interpret measure_space M by fact
-  show "finite_measure M"
-  proof
-    show "measure M (space M) \<noteq> \<infinity>" by fact
-    show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. \<mu> (A i) \<noteq> \<infinity>)"
-      using * by (auto intro!: exI[of _ "\<lambda>x. space M"])
-  qed
-qed
-
-lemma (in finite_measure) finite_measure[simp, intro]:
-  assumes "A \<in> sets M"
-  shows "\<mu> A \<noteq> \<infinity>"
-proof -
-  from `A \<in> sets M` have "A \<subseteq> space M"
-    using sets_into_space by blast
-  then have "\<mu> A \<le> \<mu> (space M)"
-    using assms top by (rule measure_mono)
-  then show ?thesis
-    using finite_measure_of_space by auto
-qed
-
-definition (in finite_measure)
-  "\<mu>' A = (if A \<in> sets M then real (\<mu> A) else 0)"
-
-lemma (in finite_measure) finite_measure_eq: "A \<in> sets M \<Longrightarrow> \<mu> A = ereal (\<mu>' A)"
-  by (auto simp: \<mu>'_def ereal_real)
-
-lemma (in finite_measure) positive_measure'[simp, intro]: "0 \<le> \<mu>' A"
-  unfolding \<mu>'_def by (auto simp: real_of_ereal_pos)
-
-lemma (in finite_measure) real_measure:
-  assumes A: "A \<in> sets M" shows "\<exists>r. 0 \<le> r \<and> \<mu> A = ereal r"
-  using finite_measure[OF A] positive_measure[OF A] by (cases "\<mu> A") auto
-
-lemma (in finite_measure) bounded_measure: "\<mu>' A \<le> \<mu>' (space M)"
-proof cases
-  assume "A \<in> sets M"
-  moreover then have "\<mu> A \<le> \<mu> (space M)"
-    using sets_into_space by (auto intro!: measure_mono)
-  ultimately show ?thesis
-    by (auto simp: \<mu>'_def intro!: real_of_ereal_positive_mono)
-qed (simp add: \<mu>'_def real_of_ereal_pos)
-
-lemma (in finite_measure) restricted_finite_measure:
-  assumes "S \<in> sets M"
-  shows "finite_measure (restricted_space S)"
-    (is "finite_measure ?r")
-proof
-  show "measure_space ?r" using restricted_measure_space[OF assms] .
-  show "measure ?r (space ?r) \<noteq> \<infinity>" 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> \<infinity>"
-  shows "finite_measure (restricted_space S)"
-proof
-  show "measure_space (restricted_space S)"
-    using `S \<in> sets M` by (rule restricted_measure_space)
-  show "measure (restricted_space S) (space (restricted_space S)) \<noteq> \<infinity>"
-    by simp fact
-qed
-
-lemma (in finite_measure) finite_measure_Diff:
-  assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
-  shows "\<mu>' (A - B) = \<mu>' A - \<mu>' B"
-  using sets[THEN finite_measure_eq]
-  using Diff[OF sets, THEN finite_measure_eq]
-  using measure_Diff[OF _ assms] by simp
-
-lemma (in finite_measure) finite_measure_Union:
-  assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
-  shows "\<mu>' (A \<union> B) = \<mu>' A + \<mu>' B"
-  using measure_additive[OF assms]
-  using sets[THEN finite_measure_eq]
-  using Un[OF sets, THEN finite_measure_eq]
-  by simp
-
-lemma (in finite_measure) finite_measure_finite_Union:
-  assumes S: "finite S" "\<And>i. i \<in> S \<Longrightarrow> A i \<in> sets M"
-  and dis: "disjoint_family_on A S"
-  shows "\<mu>' (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. \<mu>' (A i))"
-  using measure_setsum[OF assms]
-  using finite_UN[of S A, OF S, THEN finite_measure_eq]
-  using S(2)[THEN finite_measure_eq]
-  by simp
-
-lemma (in finite_measure) finite_measure_UNION:
-  assumes A: "range A \<subseteq> sets M" "disjoint_family A"
-  shows "(\<lambda>i. \<mu>' (A i)) sums (\<mu>' (\<Union>i. A i))"
-  using real_measure_UNION[OF A]
-  using countable_UN[OF A(1), THEN finite_measure_eq]
-  using A(1)[THEN subsetD, THEN finite_measure_eq]
-  by auto
-
-lemma (in finite_measure) finite_measure_mono:
-  assumes B: "B \<in> sets M" and "A \<subseteq> B" shows "\<mu>' A \<le> \<mu>' B"
-proof cases
-  assume "A \<in> sets M"
-  from this[THEN finite_measure_eq] B[THEN finite_measure_eq]
-  show ?thesis using measure_mono[OF `A \<subseteq> B` `A \<in> sets M` `B \<in> sets M`] by simp
-next
-  assume "A \<notin> sets M" then show ?thesis
-    using positive_measure'[of B] unfolding \<mu>'_def by auto
-qed
-
-lemma (in finite_measure) finite_measure_subadditive:
-  assumes m: "A \<in> sets M" "B \<in> sets M"
-  shows "\<mu>' (A \<union> B) \<le> \<mu>' A + \<mu>' B"
-  using measure_subadditive[OF m]
-  using m[THEN finite_measure_eq] Un[OF m, THEN finite_measure_eq] by simp
-
-lemma (in finite_measure) finite_measure_subadditive_finite:
-  assumes "finite I" "\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets M"
-  shows "\<mu>' (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu>' (A i))"
-  using measure_subadditive_finite[OF assms] assms
-  by (simp add: finite_measure_eq finite_UN)
-
-lemma (in finite_measure) finite_measure_countably_subadditive:
-  assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. \<mu>' (A i))"
-  shows "\<mu>' (\<Union>i. A i) \<le> (\<Sum>i. \<mu>' (A i))"
-proof -
-  note A[THEN subsetD, THEN finite_measure_eq, simp]
-  note countable_UN[OF A, THEN finite_measure_eq, simp]
-  from `summable (\<lambda>i. \<mu>' (A i))`
-  have "(\<lambda>i. ereal (\<mu>' (A i))) sums ereal (\<Sum>i. \<mu>' (A i))"
-    by (simp add: sums_ereal) (rule summable_sums)
-  from sums_unique[OF this, symmetric]
-       measure_countably_subadditive[OF A]
-  show ?thesis by simp
-qed
-
-lemma (in finite_measure) finite_measure_finite_singleton:
-  assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
-  shows "\<mu>' S = (\<Sum>x\<in>S. \<mu>' {x})"
-  using real_measure_setsum_singleton[OF assms]
-  using *[THEN finite_measure_eq]
-  using finite_UN[of S "\<lambda>x. {x}", OF assms, THEN finite_measure_eq]
-  by simp
-
-lemma (in finite_measure) finite_continuity_from_below:
-  assumes A: "range A \<subseteq> sets M" and "incseq A"
-  shows "(\<lambda>i. \<mu>' (A i)) ----> \<mu>' (\<Union>i. A i)"
-  using real_continuity_from_below[OF A, OF `incseq A` finite_measure] assms
-  using A[THEN subsetD, THEN finite_measure_eq]
-  using countable_UN[OF A, THEN finite_measure_eq]
-  by auto
-
-lemma (in finite_measure) finite_continuity_from_above:
-  assumes A: "range A \<subseteq> sets M" and "decseq A"
-  shows "(\<lambda>n. \<mu>' (A n)) ----> \<mu>' (\<Inter>i. A i)"
-  using real_continuity_from_above[OF A, OF `decseq A` finite_measure] assms
-  using A[THEN subsetD, THEN finite_measure_eq]
-  using countable_INT[OF A, THEN finite_measure_eq]
-  by auto
-
-lemma (in finite_measure) finite_measure_compl:
-  assumes S: "S \<in> sets M"
-  shows "\<mu>' (space M - S) = \<mu>' (space M) - \<mu>' S"
-  using measure_compl[OF S, OF finite_measure, OF S]
-  using S[THEN finite_measure_eq]
-  using compl_sets[OF S, THEN finite_measure_eq]
-  using top[THEN finite_measure_eq]
-  by simp
-
-lemma (in finite_measure) finite_measure_inter_full_set:
-  assumes S: "S \<in> sets M" "T \<in> sets M"
-  assumes T: "\<mu>' T = \<mu>' (space M)"
-  shows "\<mu>' (S \<inter> T) = \<mu>' S"
-  using measure_inter_full_set[OF S finite_measure]
-  using T Diff[OF S(2,1)] Diff[OF S, THEN finite_measure_eq]
-  using Int[OF S, THEN finite_measure_eq]
-  using S[THEN finite_measure_eq] top[THEN finite_measure_eq]
-  by simp
-
-lemma (in finite_measure) empty_measure'[simp]: "\<mu>' {} = 0"
-  unfolding \<mu>'_def by simp
-
-section "Finite spaces"
-
-locale finite_measure_space = finite_measure + finite_sigma_algebra
-
-lemma finite_measure_spaceI[Pure.intro!]:
-  assumes "finite (space M)"
-  assumes sets_Pow: "sets M = Pow (space M)"
-    and space: "measure M (space M) \<noteq> \<infinity>"
-    and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> measure M {x}"
-    and add: "\<And>A. A \<subseteq> space M \<Longrightarrow> measure M A = (\<Sum>x\<in>A. measure M {x})"
-  shows "finite_measure_space M"
-proof -
-  interpret finite_sigma_algebra M
-  proof
-    show "finite (space M)" by fact
-  qed (auto simp: sets_Pow)
-  interpret measure_space M
-  proof (rule finite_additivity_sufficient)
-    show "sigma_algebra M" by default
-    show "finite (space M)" by fact
-    show "positive M (measure M)"
-      by (auto simp: add positive_def intro!: setsum_nonneg pos)
-    show "additive M (measure M)"
-      using `finite (space M)`
-      by (auto simp add: additive_def add
-               intro!: setsum_Un_disjoint dest: finite_subset)
-  qed
-  interpret finite_measure M
-  proof
-    show "\<mu> (space M) \<noteq> \<infinity>" by fact
-  qed default
-  show "finite_measure_space M" 
-    by default
-qed
-
-lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)"
-  using measure_setsum[of "space M" "\<lambda>i. {i}"]
-  by (simp add: disjoint_family_on_def finite_space)
-
-lemma (in finite_measure_space) finite_measure_singleton:
-  assumes A: "A \<subseteq> space M" shows "\<mu>' A = (\<Sum>x\<in>A. \<mu>' {x})"
-  using A finite_subset[OF A finite_space]
-  by (intro finite_measure_finite_singleton) auto
-
-lemma (in finite_measure_space) finite_measure_subadditive_setsum:
-  assumes "finite I"
-  shows "\<mu>' (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu>' (A i))"
-proof cases
-  assume "(\<Union>i\<in>I. A i) \<subseteq> space M"
-  then have "\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets M" by auto
-  from finite_measure_subadditive_finite[OF `finite I` this]
-  show ?thesis by auto
-next
-  assume "\<not> (\<Union>i\<in>I. A i) \<subseteq> space M"
-  then have "\<mu>' (\<Union>i\<in>I. A i) = 0"
-    by (simp add: \<mu>'_def)
-  also have "0 \<le> (\<Sum>i\<in>I. \<mu>' (A i))"
-    by (auto intro!: setsum_nonneg)
-  finally show ?thesis .
-qed
-
-lemma suminf_cmult_indicator:
-  fixes f :: "nat \<Rightarrow> ereal"
-  assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
-  shows "(\<Sum>n. f n * indicator (A n) x) = f i"
-proof -
-  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
-    using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
-  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
-    by (auto simp: setsum_cases)
-  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
-  proof (rule ereal_SUPI)
-    fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
-    from this[of "Suc i"] show "f i \<le> y" by auto
-  qed (insert assms, simp)
-  ultimately show ?thesis using assms
-    by (subst suminf_ereal_eq_SUPR) (auto simp: indicator_def)
-qed
-
-lemma suminf_indicator:
-  assumes "disjoint_family A"
-  shows "(\<Sum>n. indicator (A n) x :: ereal) = indicator (\<Union>i. A i) x"
-proof cases
-  assume *: "x \<in> (\<Union>i. A i)"
-  then obtain i where "x \<in> A i" by auto
-  from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"]
-  show ?thesis using * by simp
-qed simp
-
-end
\ No newline at end of file
--- /dev/null	Thu Jan 01 00:00:00 1970 +0000
+++ b/src/HOL/Probability/Measure_Space.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -0,0 +1,1457 @@
+(*  Title:      HOL/Probability/Measure_Space.thy
+    Author:     Lawrence C Paulson
+    Author:     Johannes Hölzl, TU München
+    Author:     Armin Heller, TU München
+*)
+
+header {* Measure spaces and their properties *}
+
+theory Measure_Space
+imports
+  Sigma_Algebra
+  "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits"
+begin
+
+lemma suminf_eq_setsum:
+  fixes f :: "nat \<Rightarrow> 'a::{comm_monoid_add, t2_space}"
+  assumes "finite {i. f i \<noteq> 0}" (is "finite ?P")
+  shows "(\<Sum>i. f i) = (\<Sum>i | f i \<noteq> 0. f i)"
+proof cases
+  assume "?P \<noteq> {}"
+  have [dest!]: "\<And>i. Suc (Max ?P) \<le> i \<Longrightarrow> f i = 0"
+    using `finite ?P` `?P \<noteq> {}` by (auto simp: Suc_le_eq) 
+  have "(\<Sum>i. f i) = (\<Sum>i<Suc (Max ?P). f i)"
+    by (rule suminf_finite) auto
+  also have "\<dots> = (\<Sum>i | f i \<noteq> 0. f i)"
+    using `finite ?P` `?P \<noteq> {}`
+    by (intro setsum_mono_zero_right) (auto simp: less_Suc_eq_le)
+  finally show ?thesis .
+qed simp
+
+lemma suminf_cmult_indicator:
+  fixes f :: "nat \<Rightarrow> ereal"
+  assumes "disjoint_family A" "x \<in> A i" "\<And>i. 0 \<le> f i"
+  shows "(\<Sum>n. f n * indicator (A n) x) = f i"
+proof -
+  have **: "\<And>n. f n * indicator (A n) x = (if n = i then f n else 0 :: ereal)"
+    using `x \<in> A i` assms unfolding disjoint_family_on_def indicator_def by auto
+  then have "\<And>n. (\<Sum>j<n. f j * indicator (A j) x) = (if i < n then f i else 0 :: ereal)"
+    by (auto simp: setsum_cases)
+  moreover have "(SUP n. if i < n then f i else 0) = (f i :: ereal)"
+  proof (rule ereal_SUPI)
+    fix y :: ereal assume "\<And>n. n \<in> UNIV \<Longrightarrow> (if i < n then f i else 0) \<le> y"
+    from this[of "Suc i"] show "f i \<le> y" by auto
+  qed (insert assms, simp)
+  ultimately show ?thesis using assms
+    by (subst suminf_ereal_eq_SUPR) (auto simp: indicator_def)
+qed
+
+lemma suminf_indicator:
+  assumes "disjoint_family A"
+  shows "(\<Sum>n. indicator (A n) x :: ereal) = indicator (\<Union>i. A i) x"
+proof cases
+  assume *: "x \<in> (\<Union>i. A i)"
+  then obtain i where "x \<in> A i" by auto
+  from suminf_cmult_indicator[OF assms(1), OF `x \<in> A i`, of "\<lambda>k. 1"]
+  show ?thesis using * by simp
+qed simp
+
+text {*
+  The type for emeasure spaces is already defined in @{theory Sigma_Algebra}, as it is also used to
+  represent sigma algebras (with an arbitrary emeasure).
+*}
+
+section "Extend binary sets"
+
+lemma LIMSEQ_binaryset:
+  assumes f: "f {} = 0"
+  shows  "(\<lambda>n. \<Sum>i<n. f (binaryset A B i)) ----> f A + f B"
+proof -
+  have "(\<lambda>n. \<Sum>i < Suc (Suc n). f (binaryset A B i)) = (\<lambda>n. f A + f B)"
+    proof
+      fix n
+      show "(\<Sum>i < Suc (Suc n). f (binaryset A B i)) = f A + f B"
+        by (induct n)  (auto simp add: binaryset_def f)
+    qed
+  moreover
+  have "... ----> f A + f B" by (rule tendsto_const)
+  ultimately
+  have "(\<lambda>n. \<Sum>i< Suc (Suc n). f (binaryset A B i)) ----> f A + f B"
+    by metis
+  hence "(\<lambda>n. \<Sum>i< n+2. f (binaryset A B i)) ----> f A + f B"
+    by simp
+  thus ?thesis by (rule LIMSEQ_offset [where k=2])
+qed
+
+lemma binaryset_sums:
+  assumes f: "f {} = 0"
+  shows  "(\<lambda>n. f (binaryset A B n)) sums (f A + f B)"
+    by (simp add: sums_def LIMSEQ_binaryset [where f=f, OF f] atLeast0LessThan)
+
+lemma suminf_binaryset_eq:
+  fixes f :: "'a set \<Rightarrow> 'b::{comm_monoid_add, t2_space}"
+  shows "f {} = 0 \<Longrightarrow> (\<Sum>n. f (binaryset A B n)) = f A + f B"
+  by (metis binaryset_sums sums_unique)
+
+section {* Properties of a premeasure @{term \<mu>} *}
+
+text {*
+  The definitions for @{const positive} and @{const countably_additive} should be here, by they are
+  necessary to define @{typ "'a measure"} in @{theory Sigma_Algebra}.
+*}
+
+definition additive where
+  "additive M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> \<mu> (x \<union> y) = \<mu> x + \<mu> y)"
+
+definition increasing where
+  "increasing M \<mu> \<longleftrightarrow> (\<forall>x\<in>M. \<forall>y\<in>M. x \<subseteq> y \<longrightarrow> \<mu> x \<le> \<mu> y)"
+
+lemma positiveD_empty:
+  "positive M f \<Longrightarrow> f {} = 0"
+  by (auto simp add: positive_def)
+
+lemma additiveD:
+  "additive M f \<Longrightarrow> x \<inter> y = {} \<Longrightarrow> x \<in> M \<Longrightarrow> y \<in> M \<Longrightarrow> f (x \<union> y) = f x + f y"
+  by (auto simp add: additive_def)
+
+lemma increasingD:
+  "increasing M f \<Longrightarrow> x \<subseteq> y \<Longrightarrow> x\<in>M \<Longrightarrow> y\<in>M \<Longrightarrow> f x \<le> f y"
+  by (auto simp add: increasing_def)
+
+lemma countably_additiveI:
+  "(\<And>A. range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i. A i) \<in> M \<Longrightarrow> (\<Sum>i. f (A i)) = f (\<Union>i. A i))
+  \<Longrightarrow> countably_additive M f"
+  by (simp add: countably_additive_def)
+
+lemma (in ring_of_sets) disjointed_additive:
+  assumes f: "positive M f" "additive M f" and A: "range A \<subseteq> M" "incseq A"
+  shows "(\<Sum>i\<le>n. f (disjointed A i)) = f (A n)"
+proof (induct n)
+  case (Suc n)
+  then have "(\<Sum>i\<le>Suc n. f (disjointed A i)) = f (A n) + f (disjointed A (Suc n))"
+    by simp
+  also have "\<dots> = f (A n \<union> disjointed A (Suc n))"
+    using A by (subst f(2)[THEN additiveD]) (auto simp: disjointed_incseq)
+  also have "A n \<union> disjointed A (Suc n) = A (Suc n)"
+    using `incseq A` by (auto dest: incseq_SucD simp: disjointed_incseq)
+  finally show ?case .
+qed simp
+
+lemma (in ring_of_sets) additive_sum:
+  fixes A:: "'i \<Rightarrow> 'a set"
+  assumes f: "positive M f" and ad: "additive M f" and "finite S"
+      and A: "A`S \<subseteq> M"
+      and disj: "disjoint_family_on A S"
+  shows  "(\<Sum>i\<in>S. f (A i)) = f (\<Union>i\<in>S. A i)"
+using `finite S` disj A proof induct
+  case empty show ?case using f by (simp add: positive_def)
+next
+  case (insert s S)
+  then have "A s \<inter> (\<Union>i\<in>S. A i) = {}"
+    by (auto simp add: disjoint_family_on_def neq_iff)
+  moreover
+  have "A s \<in> M" using insert by blast
+  moreover have "(\<Union>i\<in>S. A i) \<in> M"
+    using insert `finite S` by auto
+  moreover
+  ultimately have "f (A s \<union> (\<Union>i\<in>S. A i)) = f (A s) + f(\<Union>i\<in>S. A i)"
+    using ad UNION_in_sets A by (auto simp add: additive_def)
+  with insert show ?case using ad disjoint_family_on_mono[of S "insert s S" A]
+    by (auto simp add: additive_def subset_insertI)
+qed
+
+lemma (in ring_of_sets) additive_increasing:
+  assumes posf: "positive M f" and addf: "additive M f"
+  shows "increasing M f"
+proof (auto simp add: increasing_def)
+  fix x y
+  assume xy: "x \<in> M" "y \<in> M" "x \<subseteq> y"
+  then have "y - x \<in> M" by auto
+  then have "0 \<le> f (y-x)" using posf[unfolded positive_def] by auto
+  then have "f x + 0 \<le> f x + f (y-x)" by (intro add_left_mono) auto
+  also have "... = f (x \<union> (y-x))" using addf
+    by (auto simp add: additive_def) (metis Diff_disjoint Un_Diff_cancel Diff xy(1,2))
+  also have "... = f y"
+    by (metis Un_Diff_cancel Un_absorb1 xy(3))
+  finally show "f x \<le> f y" by simp
+qed
+
+lemma (in ring_of_sets) countably_additive_additive:
+  assumes posf: "positive M f" and ca: "countably_additive M f"
+  shows "additive M f"
+proof (auto simp add: additive_def)
+  fix x y
+  assume x: "x \<in> M" and y: "y \<in> M" and "x \<inter> y = {}"
+  hence "disjoint_family (binaryset x y)"
+    by (auto simp add: disjoint_family_on_def binaryset_def)
+  hence "range (binaryset x y) \<subseteq> M \<longrightarrow>
+         (\<Union>i. binaryset x y i) \<in> M \<longrightarrow>
+         f (\<Union>i. binaryset x y i) = (\<Sum> n. f (binaryset x y n))"
+    using ca
+    by (simp add: countably_additive_def)
+  hence "{x,y,{}} \<subseteq> M \<longrightarrow> x \<union> y \<in> M \<longrightarrow>
+         f (x \<union> y) = (\<Sum>n. f (binaryset x y n))"
+    by (simp add: range_binaryset_eq UN_binaryset_eq)
+  thus "f (x \<union> y) = f x + f y" using posf x y
+    by (auto simp add: Un suminf_binaryset_eq positive_def)
+qed
+
+lemma (in algebra) increasing_additive_bound:
+  fixes A:: "nat \<Rightarrow> 'a set" and  f :: "'a set \<Rightarrow> ereal"
+  assumes f: "positive M f" and ad: "additive M f"
+      and inc: "increasing M f"
+      and A: "range A \<subseteq> M"
+      and disj: "disjoint_family A"
+  shows  "(\<Sum>i. f (A i)) \<le> f \<Omega>"
+proof (safe intro!: suminf_bound)
+  fix N
+  note disj_N = disjoint_family_on_mono[OF _ disj, of "{..<N}"]
+  have "(\<Sum>i<N. f (A i)) = f (\<Union>i\<in>{..<N}. A i)"
+    using A by (intro additive_sum [OF f ad _ _]) (auto simp: disj_N)
+  also have "... \<le> f \<Omega>" using space_closed A
+    by (intro increasingD[OF inc] finite_UN) auto
+  finally show "(\<Sum>i<N. f (A i)) \<le> f \<Omega>" by simp
+qed (insert f A, auto simp: positive_def)
+
+lemma (in ring_of_sets) countably_additiveI_finite:
+  assumes "finite \<Omega>" "positive M \<mu>" "additive M \<mu>"
+  shows "countably_additive M \<mu>"
+proof (rule countably_additiveI)
+  fix F :: "nat \<Rightarrow> 'a set" assume F: "range F \<subseteq> M" "(\<Union>i. F i) \<in> M" and disj: "disjoint_family F"
+
+  have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
+  from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
+
+  have inj_f: "inj_on f {i. F i \<noteq> {}}"
+  proof (rule inj_onI, simp)
+    fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
+    then have "f i \<in> F i" "f j \<in> F j" using f by force+
+    with disj * show "i = j" by (auto simp: disjoint_family_on_def)
+  qed
+  have "finite (\<Union>i. F i)"
+    by (metis F(2) assms(1) infinite_super sets_into_space)
+
+  have F_subset: "{i. \<mu> (F i) \<noteq> 0} \<subseteq> {i. F i \<noteq> {}}"
+    by (auto simp: positiveD_empty[OF `positive M \<mu>`])
+  moreover have fin_not_empty: "finite {i. F i \<noteq> {}}"
+  proof (rule finite_imageD)
+    from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
+    then show "finite (f`{i. F i \<noteq> {}})"
+      by (rule finite_subset) fact
+  qed fact
+  ultimately have fin_not_0: "finite {i. \<mu> (F i) \<noteq> 0}"
+    by (rule finite_subset)
+
+  have disj_not_empty: "disjoint_family_on F {i. F i \<noteq> {}}"
+    using disj by (auto simp: disjoint_family_on_def)
+
+  from fin_not_0 have "(\<Sum>i. \<mu> (F i)) = (\<Sum>i | \<mu> (F i) \<noteq> 0. \<mu> (F i))"
+    by (rule suminf_eq_setsum)
+  also have "\<dots> = (\<Sum>i | F i \<noteq> {}. \<mu> (F i))"
+    using fin_not_empty F_subset by (rule setsum_mono_zero_left) auto
+  also have "\<dots> = \<mu> (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
+    using `positive M \<mu>` `additive M \<mu>` fin_not_empty disj_not_empty F by (intro additive_sum) auto
+  also have "\<dots> = \<mu> (\<Union>i. F i)"
+    by (rule arg_cong[where f=\<mu>]) auto
+  finally show "(\<Sum>i. \<mu> (F i)) = \<mu> (\<Union>i. F i)" .
+qed
+
+section {* Properties of @{const emeasure} *}
+
+lemma emeasure_positive: "positive (sets M) (emeasure M)"
+  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
+
+lemma emeasure_empty[simp, intro]: "emeasure M {} = 0"
+  using emeasure_positive[of M] by (simp add: positive_def)
+
+lemma emeasure_nonneg[intro!]: "0 \<le> emeasure M A"
+  using emeasure_notin_sets[of A M] emeasure_positive[of M]
+  by (cases "A \<in> sets M") (auto simp: positive_def)
+
+lemma emeasure_not_MInf[simp]: "emeasure M A \<noteq> - \<infinity>"
+  using emeasure_nonneg[of M A] by auto
+  
+lemma emeasure_countably_additive: "countably_additive (sets M) (emeasure M)"
+  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
+
+lemma suminf_emeasure:
+  "range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
+  using countable_UN[of A UNIV M] emeasure_countably_additive[of M]
+  by (simp add: countably_additive_def)
+
+lemma emeasure_additive: "additive (sets M) (emeasure M)"
+  by (metis countably_additive_additive emeasure_positive emeasure_countably_additive)
+
+lemma plus_emeasure:
+  "a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b = {} \<Longrightarrow> emeasure M a + emeasure M b = emeasure M (a \<union> b)"
+  using additiveD[OF emeasure_additive] ..
+
+lemma setsum_emeasure:
+  "F`I \<subseteq> sets M \<Longrightarrow> disjoint_family_on F I \<Longrightarrow> finite I \<Longrightarrow>
+    (\<Sum>i\<in>I. emeasure M (F i)) = emeasure M (\<Union>i\<in>I. F i)"
+  by (metis additive_sum emeasure_positive emeasure_additive)
+
+lemma emeasure_mono:
+  "a \<subseteq> b \<Longrightarrow> b \<in> sets M \<Longrightarrow> emeasure M a \<le> emeasure M b"
+  by (metis additive_increasing emeasure_additive emeasure_nonneg emeasure_notin_sets
+            emeasure_positive increasingD)
+
+lemma emeasure_space:
+  "emeasure M A \<le> emeasure M (space M)"
+  by (metis emeasure_mono emeasure_nonneg emeasure_notin_sets sets_into_space top)
+
+lemma emeasure_compl:
+  assumes s: "s \<in> sets M" and fin: "emeasure M s \<noteq> \<infinity>"
+  shows "emeasure M (space M - s) = emeasure M (space M) - emeasure M s"
+proof -
+  from s have "0 \<le> emeasure M s" by auto
+  have "emeasure M (space M) = emeasure M (s \<union> (space M - s))" using s
+    by (metis Un_Diff_cancel Un_absorb1 s sets_into_space)
+  also have "... = emeasure M s + emeasure M (space M - s)"
+    by (rule plus_emeasure[symmetric]) (auto simp add: s)
+  finally have "emeasure M (space M) = emeasure M s + emeasure M (space M - s)" .
+  then show ?thesis
+    using fin `0 \<le> emeasure M s`
+    unfolding ereal_eq_minus_iff by (auto simp: ac_simps)
+qed
+
+lemma emeasure_Diff:
+  assumes finite: "emeasure M B \<noteq> \<infinity>"
+  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
+  shows "emeasure M (A - B) = emeasure M A - emeasure M B"
+proof -
+  have "0 \<le> emeasure M B" using assms by auto
+  have "(A - B) \<union> B = A" using `B \<subseteq> A` by auto
+  then have "emeasure M A = emeasure M ((A - B) \<union> B)" by simp
+  also have "\<dots> = emeasure M (A - B) + emeasure M B"
+    using measurable by (subst plus_emeasure[symmetric]) auto
+  finally show "emeasure M (A - B) = emeasure M A - emeasure M B"
+    unfolding ereal_eq_minus_iff
+    using finite `0 \<le> emeasure M B` by auto
+qed
+
+lemma emeasure_countable_increasing:
+  assumes A: "range A \<subseteq> sets M"
+      and A0: "A 0 = {}"
+      and ASuc: "\<And>n. A n \<subseteq> A (Suc n)"
+  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
+proof -
+  { fix n
+    have "emeasure M (A n) = (\<Sum>i<n. emeasure M (A (Suc i) - A i))"
+      proof (induct n)
+        case 0 thus ?case by (auto simp add: A0)
+      next
+        case (Suc m)
+        have "A (Suc m) = A m \<union> (A (Suc m) - A m)"
+          by (metis ASuc Un_Diff_cancel Un_absorb1)
+        hence "emeasure M (A (Suc m)) =
+               emeasure M (A m) + emeasure M (A (Suc m) - A m)"
+          by (subst plus_emeasure)
+             (auto simp add: emeasure_additive range_subsetD [OF A])
+        with Suc show ?case
+          by simp
+      qed }
+  note Meq = this
+  have Aeq: "(\<Union>i. A (Suc i) - A i) = (\<Union>i. A i)"
+    proof (rule UN_finite2_eq [where k=1], simp)
+      fix i
+      show "(\<Union>i\<in>{0..<i}. A (Suc i) - A i) = (\<Union>i\<in>{0..<Suc i}. A i)"
+        proof (induct i)
+          case 0 thus ?case by (simp add: A0)
+        next
+          case (Suc i)
+          thus ?case
+            by (auto simp add: atLeastLessThanSuc intro: subsetD [OF ASuc])
+        qed
+    qed
+  have A1: "\<And>i. A (Suc i) - A i \<in> sets M"
+    by (metis A Diff range_subsetD)
+  have A2: "(\<Union>i. A (Suc i) - A i) \<in> sets M"
+    by (blast intro: range_subsetD [OF A])
+  have "(SUP n. \<Sum>i<n. emeasure M (A (Suc i) - A i)) = (\<Sum>i. emeasure M (A (Suc i) - A i))"
+    using A by (auto intro!: suminf_ereal_eq_SUPR[symmetric])
+  also have "\<dots> = emeasure M (\<Union>i. A (Suc i) - A i)"
+    by (rule suminf_emeasure)
+       (auto simp add: disjoint_family_Suc ASuc A1 A2)
+  also have "... =  emeasure M (\<Union>i. A i)"
+    by (simp add: Aeq)
+  finally have "(SUP n. \<Sum>i<n. emeasure M (A (Suc i) - A i)) = emeasure M (\<Union>i. A i)" .
+  then show ?thesis by (auto simp add: Meq)
+qed
+
+lemma SUP_emeasure_incseq:
+  assumes A: "range A \<subseteq> sets M" and "incseq A"
+  shows "(SUP n. emeasure M (A n)) = emeasure M (\<Union>i. A i)"
+proof -
+  have *: "(SUP n. emeasure M (nat_case {} A (Suc n))) = (SUP n. emeasure M (nat_case {} A n))"
+    using A by (auto intro!: SUPR_eq exI split: nat.split)
+  have ueq: "(\<Union>i. nat_case {} A i) = (\<Union>i. A i)"
+    by (auto simp add: split: nat.splits)
+  have meq: "\<And>n. emeasure M (A n) = (emeasure M \<circ> nat_case {} A) (Suc n)"
+    by simp
+  have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\<Union>i. nat_case {} A i)"
+    using range_subsetD[OF A] incseq_SucD[OF `incseq A`]
+    by (force split: nat.splits intro!: emeasure_countable_increasing)
+  also have "emeasure M (\<Union>i. nat_case {} A i) = emeasure M (\<Union>i. A i)"
+    by (simp add: ueq)
+  finally have "(SUP n. emeasure M (nat_case {} A n)) = emeasure M (\<Union>i. A i)" .
+  thus ?thesis unfolding meq * comp_def .
+qed
+
+lemma incseq_emeasure:
+  assumes "range B \<subseteq> sets M" "incseq B"
+  shows "incseq (\<lambda>i. emeasure M (B i))"
+  using assms by (auto simp: incseq_def intro!: emeasure_mono)
+
+lemma Lim_emeasure_incseq:
+  assumes A: "range A \<subseteq> sets M" "incseq A"
+  shows "(\<lambda>i. (emeasure M (A i))) ----> emeasure M (\<Union>i. A i)"
+  using LIMSEQ_ereal_SUPR[OF incseq_emeasure, OF A]
+    SUP_emeasure_incseq[OF A] by simp
+
+lemma decseq_emeasure:
+  assumes "range B \<subseteq> sets M" "decseq B"
+  shows "decseq (\<lambda>i. emeasure M (B i))"
+  using assms by (auto simp: decseq_def intro!: emeasure_mono)
+
+lemma INF_emeasure_decseq:
+  assumes A: "range A \<subseteq> sets M" and "decseq A"
+  and finite: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+  shows "(INF n. emeasure M (A n)) = emeasure M (\<Inter>i. A i)"
+proof -
+  have le_MI: "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
+    using A by (auto intro!: emeasure_mono)
+  hence *: "emeasure M (\<Inter>i. A i) \<noteq> \<infinity>" using finite[of 0] by auto
+
+  have A0: "0 \<le> emeasure M (A 0)" using A by auto
+
+  have "emeasure M (A 0) - (INF n. emeasure M (A n)) = emeasure M (A 0) + (SUP n. - emeasure M (A n))"
+    by (simp add: ereal_SUPR_uminus minus_ereal_def)
+  also have "\<dots> = (SUP n. emeasure M (A 0) - emeasure M (A n))"
+    unfolding minus_ereal_def using A0 assms
+    by (subst SUPR_ereal_add) (auto simp add: decseq_emeasure)
+  also have "\<dots> = (SUP n. emeasure M (A 0 - A n))"
+    using A finite `decseq A`[unfolded decseq_def] by (subst emeasure_Diff) auto
+  also have "\<dots> = emeasure M (\<Union>i. A 0 - A i)"
+  proof (rule SUP_emeasure_incseq)
+    show "range (\<lambda>n. A 0 - A n) \<subseteq> sets M"
+      using A by auto
+    show "incseq (\<lambda>n. A 0 - A n)"
+      using `decseq A` by (auto simp add: incseq_def decseq_def)
+  qed
+  also have "\<dots> = emeasure M (A 0) - emeasure M (\<Inter>i. A i)"
+    using A finite * by (simp, subst emeasure_Diff) auto
+  finally show ?thesis
+    unfolding ereal_minus_eq_minus_iff using finite A0 by auto
+qed
+
+lemma Lim_emeasure_decseq:
+  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+  shows "(\<lambda>i. emeasure M (A i)) ----> emeasure M (\<Inter>i. A i)"
+  using LIMSEQ_ereal_INFI[OF decseq_emeasure, OF A]
+  using INF_emeasure_decseq[OF A fin] by simp
+
+lemma emeasure_subadditive:
+  assumes measurable: "A \<in> sets M" "B \<in> sets M"
+  shows "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
+proof -
+  from plus_emeasure[of A M "B - A"]
+  have "emeasure M (A \<union> B) = emeasure M A + emeasure M (B - A)"
+    using assms by (simp add: Diff)
+  also have "\<dots> \<le> emeasure M A + emeasure M B"
+    using assms by (auto intro!: add_left_mono emeasure_mono)
+  finally show ?thesis .
+qed
+
+lemma emeasure_subadditive_finite:
+  assumes "finite I" "A ` I \<subseteq> sets M"
+  shows "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
+using assms proof induct
+  case (insert i I)
+  then have "emeasure M (\<Union>i\<in>insert i I. A i) = emeasure M (A i \<union> (\<Union>i\<in>I. A i))"
+    by simp
+  also have "\<dots> \<le> emeasure M (A i) + emeasure M (\<Union>i\<in>I. A i)"
+    using insert by (intro emeasure_subadditive finite_UN) auto
+  also have "\<dots> \<le> emeasure M (A i) + (\<Sum>i\<in>I. emeasure M (A i))"
+    using insert by (intro add_mono) auto
+  also have "\<dots> = (\<Sum>i\<in>insert i I. emeasure M (A i))"
+    using insert by auto
+  finally show ?case .
+qed simp
+
+lemma emeasure_subadditive_countably:
+  assumes "range f \<subseteq> sets M"
+  shows "emeasure M (\<Union>i. f i) \<le> (\<Sum>i. emeasure M (f i))"
+proof -
+  have "emeasure M (\<Union>i. f i) = emeasure M (\<Union>i. disjointed f i)"
+    unfolding UN_disjointed_eq ..
+  also have "\<dots> = (\<Sum>i. emeasure M (disjointed f i))"
+    using range_disjointed_sets[OF assms] suminf_emeasure[of "disjointed f"]
+    by (simp add:  disjoint_family_disjointed comp_def)
+  also have "\<dots> \<le> (\<Sum>i. emeasure M (f i))"
+    using range_disjointed_sets[OF assms] assms
+    by (auto intro!: suminf_le_pos emeasure_mono disjointed_subset)
+  finally show ?thesis .
+qed
+
+lemma emeasure_insert:
+  assumes sets: "{x} \<in> sets M" "A \<in> sets M" and "x \<notin> A"
+  shows "emeasure M (insert x A) = emeasure M {x} + emeasure M A"
+proof -
+  have "{x} \<inter> A = {}" using `x \<notin> A` by auto
+  from plus_emeasure[OF sets this] show ?thesis by simp
+qed
+
+lemma emeasure_eq_setsum_singleton:
+  assumes "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+  shows "emeasure M S = (\<Sum>x\<in>S. emeasure M {x})"
+  using setsum_emeasure[of "\<lambda>x. {x}" S M] assms
+  by (auto simp: disjoint_family_on_def subset_eq)
+
+lemma setsum_emeasure_cover:
+  assumes "finite S" and "A \<in> sets M" and br_in_M: "B ` S \<subseteq> sets M"
+  assumes A: "A \<subseteq> (\<Union>i\<in>S. B i)"
+  assumes disj: "disjoint_family_on B S"
+  shows "emeasure M A = (\<Sum>i\<in>S. emeasure M (A \<inter> (B i)))"
+proof -
+  have "(\<Sum>i\<in>S. emeasure M (A \<inter> (B i))) = emeasure M (\<Union>i\<in>S. A \<inter> (B i))"
+  proof (rule setsum_emeasure)
+    show "disjoint_family_on (\<lambda>i. A \<inter> B i) S"
+      using `disjoint_family_on B S`
+      unfolding disjoint_family_on_def by auto
+  qed (insert assms, auto)
+  also have "(\<Union>i\<in>S. A \<inter> (B i)) = A"
+    using A by auto
+  finally show ?thesis by simp
+qed
+
+lemma emeasure_eq_0:
+  "N \<in> sets M \<Longrightarrow> emeasure M N = 0 \<Longrightarrow> K \<subseteq> N \<Longrightarrow> emeasure M K = 0"
+  by (metis emeasure_mono emeasure_nonneg order_eq_iff)
+
+lemma emeasure_UN_eq_0:
+  assumes "\<And>i::nat. emeasure M (N i) = 0" and "range N \<subseteq> sets M"
+  shows "emeasure M (\<Union> i. N i) = 0"
+proof -
+  have "0 \<le> emeasure M (\<Union> i. N i)" using assms by auto
+  moreover have "emeasure M (\<Union> i. N i) \<le> 0"
+    using emeasure_subadditive_countably[OF assms(2)] assms(1) by simp
+  ultimately show ?thesis by simp
+qed
+
+lemma measure_eqI_finite:
+  assumes [simp]: "sets M = Pow A" "sets N = Pow A" and "finite A"
+  assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
+  shows "M = N"
+proof (rule measure_eqI)
+  fix X assume "X \<in> sets M"
+  then have X: "X \<subseteq> A" by auto
+  then have "emeasure M X = (\<Sum>a\<in>X. emeasure M {a})"
+    using `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
+  also have "\<dots> = (\<Sum>a\<in>X. emeasure N {a})"
+    using X eq by (auto intro!: setsum_cong)
+  also have "\<dots> = emeasure N X"
+    using X `finite A` by (subst emeasure_eq_setsum_singleton) (auto dest: finite_subset)
+  finally show "emeasure M X = emeasure N X" .
+qed simp
+
+lemma measure_eqI_generator_eq:
+  fixes M N :: "'a measure" and E :: "'a set set" and A :: "nat \<Rightarrow> 'a set"
+  assumes "Int_stable E" "E \<subseteq> Pow \<Omega>"
+  and eq: "\<And>X. X \<in> E \<Longrightarrow> emeasure M X = emeasure N X"
+  and M: "sets M = sigma_sets \<Omega> E"
+  and N: "sets N = sigma_sets \<Omega> E"
+  and A: "range A \<subseteq> E" "incseq A" "(\<Union>i. A i) = \<Omega>" "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+  shows "M = N"
+proof -
+  let ?D = "\<lambda>F. {D. D \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> D) = emeasure N (F \<inter> D)}"
+  interpret S: sigma_algebra \<Omega> "sigma_sets \<Omega> E" by (rule sigma_algebra_sigma_sets) fact
+  { fix F assume "F \<in> E" and "emeasure M F \<noteq> \<infinity>"
+    then have [intro]: "F \<in> sigma_sets \<Omega> E" by auto
+    have "emeasure N F \<noteq> \<infinity>" using `emeasure M F \<noteq> \<infinity>` `F \<in> E` eq by simp
+    interpret D: dynkin_system \<Omega> "?D F"
+    proof (rule dynkin_systemI, simp_all)
+      fix A assume "A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> A) = emeasure N (F \<inter> A)"
+      then show "A \<subseteq> \<Omega>" using S.sets_into_space by auto
+    next
+      have "F \<inter> \<Omega> = F" using `F \<in> E` `E \<subseteq> Pow \<Omega>` by auto
+      then show "emeasure M (F \<inter> \<Omega>) = emeasure N (F \<inter> \<Omega>)"
+        using `F \<in> E` eq by (auto intro: sigma_sets_top)
+    next
+      fix A assume *: "A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> A) = emeasure N (F \<inter> A)"
+      then have **: "F \<inter> (\<Omega> - A) = F - (F \<inter> A)"
+        and [intro]: "F \<inter> A \<in> sigma_sets \<Omega> E"
+        using `F \<in> E` S.sets_into_space by auto
+      have "emeasure N (F \<inter> A) \<le> emeasure N F" by (auto intro!: emeasure_mono simp: M N)
+      then have "emeasure N (F \<inter> A) \<noteq> \<infinity>" using `emeasure N F \<noteq> \<infinity>` by auto
+      have "emeasure M (F \<inter> A) \<le> emeasure M F" by (auto intro!: emeasure_mono simp: M N)
+      then have "emeasure M (F \<inter> A) \<noteq> \<infinity>" using `emeasure M F \<noteq> \<infinity>` by auto
+      then have "emeasure M (F \<inter> (\<Omega> - A)) = emeasure M F - emeasure M (F \<inter> A)" unfolding **
+        using `F \<inter> A \<in> sigma_sets \<Omega> E` by (auto intro!: emeasure_Diff simp: M N)
+      also have "\<dots> = emeasure N F - emeasure N (F \<inter> A)" using eq `F \<in> E` * by simp
+      also have "\<dots> = emeasure N (F \<inter> (\<Omega> - A))" unfolding **
+        using `F \<inter> A \<in> sigma_sets \<Omega> E` `emeasure N (F \<inter> A) \<noteq> \<infinity>`
+        by (auto intro!: emeasure_Diff[symmetric] simp: M N)
+      finally show "\<Omega> - A \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> (\<Omega> - A)) = emeasure N (F \<inter> (\<Omega> - A))"
+        using * by auto
+    next
+      fix A :: "nat \<Rightarrow> 'a set"
+      assume "disjoint_family A" "range A \<subseteq> {X \<in> sigma_sets \<Omega> E. emeasure M (F \<inter> X) = emeasure N (F \<inter> X)}"
+      then have A: "range (\<lambda>i. F \<inter> A i) \<subseteq> sigma_sets \<Omega> E" "F \<inter> (\<Union>x. A x) = (\<Union>x. F \<inter> A x)"
+        "disjoint_family (\<lambda>i. F \<inter> A i)" "\<And>i. emeasure M (F \<inter> A i) = emeasure N (F \<inter> A i)" "range A \<subseteq> sigma_sets \<Omega> E"
+        by (auto simp: disjoint_family_on_def subset_eq)
+      then show "(\<Union>x. A x) \<in> sigma_sets \<Omega> E \<and> emeasure M (F \<inter> (\<Union>x. A x)) = emeasure N (F \<inter> (\<Union>x. A x))"
+        by (auto simp: M N suminf_emeasure[symmetric] simp del: UN_simps)
+    qed
+    have *: "sigma_sets \<Omega> E = ?D F"
+      using `F \<in> E` `Int_stable E`
+      by (intro D.dynkin_lemma) (auto simp add: Int_stable_def eq)
+    have "\<And>D. D \<in> sigma_sets \<Omega> E \<Longrightarrow> emeasure M (F \<inter> D) = emeasure N (F \<inter> D)"
+      by (subst (asm) *) auto }
+  note * = this
+  show "M = N"
+  proof (rule measure_eqI)
+    show "sets M = sets N"
+      using M N by simp
+    fix X assume "X \<in> sets M"
+    then have "X \<in> sigma_sets \<Omega> E"
+      using M by simp
+    let ?A = "\<lambda>i. A i \<inter> X"
+    have "range ?A \<subseteq> sigma_sets \<Omega> E" "incseq ?A"
+      using A(1,2) `X \<in> sigma_sets \<Omega> E` by (auto simp: incseq_def)
+    moreover
+    { fix i have "emeasure M (?A i) = emeasure N (?A i)"
+        using *[of "A i" X] `X \<in> sigma_sets \<Omega> E` A finite by auto }
+    ultimately show "emeasure M X = emeasure N X"
+      using SUP_emeasure_incseq[of ?A M] SUP_emeasure_incseq[of ?A N] A(3) `X \<in> sigma_sets \<Omega> E`
+      by (auto simp: M N SUP_emeasure_incseq)
+  qed
+qed
+
+lemma measure_of_of_measure: "measure_of (space M) (sets M) (emeasure M) = M"
+proof (intro measure_eqI emeasure_measure_of_sigma)
+  show "sigma_algebra (space M) (sets M)" ..
+  show "positive (sets M) (emeasure M)"
+    by (simp add: positive_def emeasure_nonneg)
+  show "countably_additive (sets M) (emeasure M)"
+    by (simp add: emeasure_countably_additive)
+qed simp_all
+
+section "@{text \<mu>}-null sets"
+
+definition null_sets :: "'a measure \<Rightarrow> 'a set set" where
+  "null_sets M = {N\<in>sets M. emeasure M N = 0}"
+
+lemma null_setsD1[dest]: "A \<in> null_sets M \<Longrightarrow> emeasure M A = 0"
+  by (simp add: null_sets_def)
+
+lemma null_setsD2[dest]: "A \<in> null_sets M \<Longrightarrow> A \<in> sets M"
+  unfolding null_sets_def by simp
+
+lemma null_setsI[intro]: "emeasure M A = 0 \<Longrightarrow> A \<in> sets M \<Longrightarrow> A \<in> null_sets M"
+  unfolding null_sets_def by simp
+
+interpretation null_sets: ring_of_sets "space M" "null_sets M" for M
+proof
+  show "null_sets M \<subseteq> Pow (space M)"
+    using sets_into_space by auto
+  show "{} \<in> null_sets M"
+    by auto
+  fix A B assume sets: "A \<in> null_sets M" "B \<in> null_sets M"
+  then have "A \<in> sets M" "B \<in> sets M"
+    by auto
+  moreover then have "emeasure M (A \<union> B) \<le> emeasure M A + emeasure M B"
+    "emeasure M (A - B) \<le> emeasure M A"
+    by (auto intro!: emeasure_subadditive emeasure_mono)
+  moreover have "emeasure M B = 0" "emeasure M A = 0"
+    using sets by auto
+  ultimately show "A - B \<in> null_sets M" "A \<union> B \<in> null_sets M"
+    by (auto intro!: antisym)
+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 SUP_def image_compose
+    unfolding surj_from_nat ..
+  then show ?thesis by simp
+qed
+
+lemma null_sets_UN[intro]:
+  assumes "\<And>i::'i::countable. N i \<in> null_sets M"
+  shows "(\<Union>i. N i) \<in> null_sets M"
+proof (intro conjI CollectI null_setsI)
+  show "(\<Union>i. N i) \<in> sets M" using assms by auto
+  have "0 \<le> emeasure M (\<Union>i. N i)" by (rule emeasure_nonneg)
+  moreover have "emeasure M (\<Union>i. N i) \<le> (\<Sum>n. emeasure M (N (Countable.from_nat n)))"
+    unfolding UN_from_nat[of N]
+    using assms by (intro emeasure_subadditive_countably) auto
+  ultimately show "emeasure M (\<Union>i. N i) = 0"
+    using assms by (auto simp: null_setsD1)
+qed
+
+lemma null_set_Int1:
+  assumes "B \<in> null_sets M" "A \<in> sets M" shows "A \<inter> B \<in> null_sets M"
+proof (intro CollectI conjI null_setsI)
+  show "emeasure M (A \<inter> B) = 0" using assms
+    by (intro emeasure_eq_0[of B _ "A \<inter> B"]) auto
+qed (insert assms, auto)
+
+lemma null_set_Int2:
+  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B \<inter> A \<in> null_sets M"
+  using assms by (subst Int_commute) (rule null_set_Int1)
+
+lemma emeasure_Diff_null_set:
+  assumes "B \<in> null_sets M" "A \<in> sets M"
+  shows "emeasure M (A - B) = emeasure M A"
+proof -
+  have *: "A - B = (A - (A \<inter> B))" by auto
+  have "A \<inter> B \<in> null_sets M" using assms by (rule null_set_Int1)
+  then show ?thesis
+    unfolding * using assms
+    by (subst emeasure_Diff) auto
+qed
+
+lemma null_set_Diff:
+  assumes "B \<in> null_sets M" "A \<in> sets M" shows "B - A \<in> null_sets M"
+proof (intro CollectI conjI null_setsI)
+  show "emeasure M (B - A) = 0" using assms by (intro emeasure_eq_0[of B _ "B - A"]) auto
+qed (insert assms, auto)
+
+lemma emeasure_Un_null_set:
+  assumes "A \<in> sets M" "B \<in> null_sets M"
+  shows "emeasure M (A \<union> B) = emeasure M A"
+proof -
+  have *: "A \<union> B = A \<union> (B - A)" by auto
+  have "B - A \<in> null_sets M" using assms(2,1) by (rule null_set_Diff)
+  then show ?thesis
+    unfolding * using assms
+    by (subst plus_emeasure[symmetric]) auto
+qed
+
+section "Formalize almost everywhere"
+
+definition ae_filter :: "'a measure \<Rightarrow> 'a filter" where
+  "ae_filter M = Abs_filter (\<lambda>P. \<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N)"
+
+abbreviation
+  almost_everywhere :: "'a measure \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> bool" where
+  "almost_everywhere M P \<equiv> eventually P (ae_filter M)"
+
+syntax
+  "_almost_everywhere" :: "pttrn \<Rightarrow> 'a \<Rightarrow> bool \<Rightarrow> bool" ("AE _ in _. _" [0,0,10] 10)
+
+translations
+  "AE x in M. P" == "CONST almost_everywhere M (%x. P)"
+
+lemma eventually_ae_filter:
+  fixes M P
+  defines [simp]: "F \<equiv> \<lambda>P. \<exists>N\<in>null_sets M. {x \<in> space M. \<not> P x} \<subseteq> N" 
+  shows "eventually P (ae_filter M) \<longleftrightarrow> F P"
+  unfolding ae_filter_def F_def[symmetric]
+proof (rule eventually_Abs_filter)
+  show "is_filter F"
+  proof
+    fix P Q assume "F P" "F Q"
+    then obtain N L where N: "N \<in> null_sets M" "{x \<in> space M. \<not> P x} \<subseteq> N"
+      and L: "L \<in> null_sets M" "{x \<in> space M. \<not> Q x} \<subseteq> L"
+      by auto
+    then have "L \<union> N \<in> null_sets M" "{x \<in> space M. \<not> (P x \<and> Q x)} \<subseteq> L \<union> N" by auto
+    then show "F (\<lambda>x. P x \<and> Q x)" by auto
+  next
+    fix P Q assume "F P"
+    then obtain N where N: "N \<in> null_sets M" "{x \<in> space M. \<not> P x} \<subseteq> N" by auto
+    moreover assume "\<forall>x. P x \<longrightarrow> Q x"
+    ultimately have "N \<in> null_sets M" "{x \<in> space M. \<not> Q x} \<subseteq> N" by auto
+    then show "F Q" by auto
+  qed auto
+qed
+
+lemma AE_I':
+  "N \<in> null_sets M \<Longrightarrow> {x\<in>space M. \<not> P x} \<subseteq> N \<Longrightarrow> (AE x in M. P x)"
+  unfolding eventually_ae_filter by auto
+
+lemma AE_iff_null:
+  assumes "{x\<in>space M. \<not> P x} \<in> sets M" (is "?P \<in> sets M")
+  shows "(AE x in M. P x) \<longleftrightarrow> {x\<in>space M. \<not> P x} \<in> null_sets M"
+proof
+  assume "AE x in M. P x" then obtain N where N: "N \<in> sets M" "?P \<subseteq> N" "emeasure M N = 0"
+    unfolding eventually_ae_filter by auto
+  have "0 \<le> emeasure M ?P" by auto
+  moreover have "emeasure M ?P \<le> emeasure M N"
+    using assms N(1,2) by (auto intro: emeasure_mono)
+  ultimately have "emeasure M ?P = 0" unfolding `emeasure M N = 0` by auto
+  then show "?P \<in> null_sets M" using assms by auto
+next
+  assume "?P \<in> null_sets M" with assms show "AE x in M. P x" by (auto intro: AE_I')
+qed
+
+lemma AE_iff_null_sets:
+  "N \<in> sets M \<Longrightarrow> N \<in> null_sets M \<longleftrightarrow> (AE x in M. x \<notin> N)"
+  using Int_absorb1[OF sets_into_space, of N M]
+  by (subst AE_iff_null) (auto simp: Int_def[symmetric])
+
+lemma AE_iff_measurable:
+  "N \<in> sets M \<Longrightarrow> {x\<in>space M. \<not> P x} = N \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> emeasure M N = 0"
+  using AE_iff_null[of _ P] by auto
+
+lemma AE_E[consumes 1]:
+  assumes "AE x in M. P x"
+  obtains N where "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
+  using assms unfolding eventually_ae_filter by auto
+
+lemma AE_E2:
+  assumes "AE x in M. P x" "{x\<in>space M. P x} \<in> sets M"
+  shows "emeasure M {x\<in>space M. \<not> P x} = 0" (is "emeasure M ?P = 0")
+proof -
+  have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" by auto
+  with AE_iff_null[of M P] assms show ?thesis by auto
+qed
+
+lemma AE_I:
+  assumes "{x \<in> space M. \<not> P x} \<subseteq> N" "emeasure M N = 0" "N \<in> sets M"
+  shows "AE x in M. P x"
+  using assms unfolding eventually_ae_filter by auto
+
+lemma AE_mp[elim!]:
+  assumes AE_P: "AE x in M. P x" and AE_imp: "AE x in M. P x \<longrightarrow> Q x"
+  shows "AE x in M. Q x"
+proof -
+  from AE_P obtain A where P: "{x\<in>space M. \<not> P x} \<subseteq> A"
+    and A: "A \<in> sets M" "emeasure M A = 0"
+    by (auto elim!: AE_E)
+
+  from AE_imp obtain B where imp: "{x\<in>space M. P x \<and> \<not> Q x} \<subseteq> B"
+    and B: "B \<in> sets M" "emeasure M B = 0"
+    by (auto elim!: AE_E)
+
+  show ?thesis
+  proof (intro AE_I)
+    have "0 \<le> emeasure M (A \<union> B)" using A B by auto
+    moreover have "emeasure M (A \<union> B) \<le> 0"
+      using emeasure_subadditive[of A M B] A B by auto
+    ultimately show "A \<union> B \<in> sets M" "emeasure M (A \<union> B) = 0" using A B by auto
+    show "{x\<in>space M. \<not> Q x} \<subseteq> A \<union> B"
+      using P imp by auto
+  qed
+qed
+
+(* depricated replace by laws about eventually *)
+lemma
+  shows AE_iffI: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<longleftrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
+    and AE_disjI1: "AE x in M. P x \<Longrightarrow> AE x in M. P x \<or> Q x"
+    and AE_disjI2: "AE x in M. Q x \<Longrightarrow> AE x in M. P x \<or> Q x"
+    and AE_conjI: "AE x in M. P x \<Longrightarrow> AE x in M. Q x \<Longrightarrow> AE x in M. P x \<and> Q x"
+    and AE_conj_iff[simp]: "(AE x in M. P x \<and> Q x) \<longleftrightarrow> (AE x in M. P x) \<and> (AE x in M. Q x)"
+  by auto
+
+lemma AE_impI:
+  "(P \<Longrightarrow> AE x in M. Q x) \<Longrightarrow> AE x in M. P \<longrightarrow> Q x"
+  by (cases P) auto
+
+lemma AE_measure:
+  assumes AE: "AE x in M. P x" and sets: "{x\<in>space M. P x} \<in> sets M" (is "?P \<in> sets M")
+  shows "emeasure M {x\<in>space M. P x} = emeasure M (space M)"
+proof -
+  from AE_E[OF AE] guess N . note N = this
+  with sets have "emeasure M (space M) \<le> emeasure M (?P \<union> N)"
+    by (intro emeasure_mono) auto
+  also have "\<dots> \<le> emeasure M ?P + emeasure M N"
+    using sets N by (intro emeasure_subadditive) auto
+  also have "\<dots> = emeasure M ?P" using N by simp
+  finally show "emeasure M ?P = emeasure M (space M)"
+    using emeasure_space[of M "?P"] by auto
+qed
+
+lemma AE_space: "AE x in M. x \<in> space M"
+  by (rule AE_I[where N="{}"]) auto
+
+lemma AE_I2[simp, intro]:
+  "(\<And>x. x \<in> space M \<Longrightarrow> P x) \<Longrightarrow> AE x in M. P x"
+  using AE_space by force
+
+lemma AE_Ball_mp:
+  "\<forall>x\<in>space M. P x \<Longrightarrow> AE x in M. P x \<longrightarrow> Q x \<Longrightarrow> AE x in M. Q x"
+  by auto
+
+lemma AE_cong[cong]:
+  "(\<And>x. x \<in> space M \<Longrightarrow> P x \<longleftrightarrow> Q x) \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> (AE x in M. Q x)"
+  by auto
+
+lemma AE_all_countable:
+  "(AE x in M. \<forall>i. P i x) \<longleftrightarrow> (\<forall>i::'i::countable. AE x in M. P i x)"
+proof
+  assume "\<forall>i. AE x in M. P i x"
+  from this[unfolded eventually_ae_filter Bex_def, THEN choice]
+  obtain N where N: "\<And>i. N i \<in> null_sets M" "\<And>i. {x\<in>space M. \<not> P i x} \<subseteq> N i" by auto
+  have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. {x\<in>space M. \<not> P i x})" by auto
+  also have "\<dots> \<subseteq> (\<Union>i. N i)" using N by auto
+  finally have "{x\<in>space M. \<not> (\<forall>i. P i x)} \<subseteq> (\<Union>i. N i)" .
+  moreover from N have "(\<Union>i. N i) \<in> null_sets M"
+    by (intro null_sets_UN) auto
+  ultimately show "AE x in M. \<forall>i. P i x"
+    unfolding eventually_ae_filter by auto
+qed auto
+
+lemma AE_finite_all:
+  assumes f: "finite S" shows "(AE x in M. \<forall>i\<in>S. P i x) \<longleftrightarrow> (\<forall>i\<in>S. AE x in M. P i x)"
+  using f by induct auto
+
+lemma AE_finite_allI:
+  assumes "finite S"
+  shows "(\<And>s. s \<in> S \<Longrightarrow> AE x in M. Q s x) \<Longrightarrow> AE x in M. \<forall>s\<in>S. Q s x"
+  using AE_finite_all[OF `finite S`] by auto
+
+lemma emeasure_mono_AE:
+  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B"
+    and B: "B \<in> sets M"
+  shows "emeasure M A \<le> emeasure M B"
+proof cases
+  assume A: "A \<in> sets M"
+  from imp obtain N where N: "{x\<in>space M. \<not> (x \<in> A \<longrightarrow> x \<in> B)} \<subseteq> N" "N \<in> null_sets M"
+    by (auto simp: eventually_ae_filter)
+  have "emeasure M A = emeasure M (A - N)"
+    using N A by (subst emeasure_Diff_null_set) auto
+  also have "emeasure M (A - N) \<le> emeasure M (B - N)"
+    using N A B sets_into_space by (auto intro!: emeasure_mono)
+  also have "emeasure M (B - N) = emeasure M B"
+    using N B by (subst emeasure_Diff_null_set) auto
+  finally show ?thesis .
+qed (simp add: emeasure_nonneg emeasure_notin_sets)
+
+lemma emeasure_eq_AE:
+  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
+  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
+  shows "emeasure M A = emeasure M B"
+  using assms by (safe intro!: antisym emeasure_mono_AE) auto
+
+section {* @{text \<sigma>}-finite Measures *}
+
+locale sigma_finite_measure =
+  fixes M :: "'a measure"
+  assumes sigma_finite: "\<exists>A::nat \<Rightarrow> 'a set.
+    range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. emeasure M (A i) \<noteq> \<infinity>)"
+
+lemma (in sigma_finite_measure) sigma_finite_disjoint:
+  obtains A :: "nat \<Rightarrow> 'a set"
+  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "disjoint_family A"
+proof atomize_elim
+  case goal1
+  obtain A :: "nat \<Rightarrow> 'a set" where
+    range: "range A \<subseteq> sets M" and
+    space: "(\<Union>i. A i) = space M" and
+    measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+    using sigma_finite by auto
+  note range' = range_disjointed_sets[OF range] range
+  { fix i
+    have "emeasure M (disjointed A i) \<le> emeasure M (A i)"
+      using range' disjointed_subset[of A i] by (auto intro!: emeasure_mono)
+    then have "emeasure M (disjointed A i) \<noteq> \<infinity>"
+      using measure[of i] by auto }
+  with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
+  show ?case by (auto intro!: exI[of _ "disjointed A"])
+qed
+
+lemma (in sigma_finite_measure) sigma_finite_incseq:
+  obtains A :: "nat \<Rightarrow> 'a set"
+  where "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And>i. emeasure M (A i) \<noteq> \<infinity>" "incseq A"
+proof atomize_elim
+  case goal1
+  obtain F :: "nat \<Rightarrow> 'a set" where
+    F: "range F \<subseteq> sets M" "(\<Union>i. F i) = space M" "\<And>i. emeasure M (F i) \<noteq> \<infinity>"
+    using sigma_finite by auto
+  then show ?case
+  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 fastforce
+  next
+    fix n
+    have "emeasure M (\<Union> i\<le>n. F i) \<le> (\<Sum>i\<le>n. emeasure M (F i))" using F
+      by (auto intro!: emeasure_subadditive_finite)
+    also have "\<dots> < \<infinity>"
+      using F by (auto simp: setsum_Pinfty)
+    finally show "emeasure M (\<Union> i\<le>n. F i) \<noteq> \<infinity>" by simp
+  qed (force simp: incseq_def)+
+qed
+
+section {* Measure space induced by distribution of @{const measurable}-functions *}
+
+definition distr :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) \<Rightarrow> 'b measure" where
+  "distr M N f = measure_of (space N) (sets N) (\<lambda>A. emeasure M (f -` A \<inter> space M))"
+
+lemma
+  shows sets_distr[simp]: "sets (distr M N f) = sets N"
+    and space_distr[simp]: "space (distr M N f) = space N"
+  by (auto simp: distr_def)
+
+lemma
+  shows measurable_distr_eq1[simp]: "measurable (distr Mf Nf f) Mf' = measurable Nf Mf'"
+    and measurable_distr_eq2[simp]: "measurable Mg' (distr Mg Ng g) = measurable Mg' Ng"
+  by (auto simp: measurable_def)
+
+lemma emeasure_distr:
+  fixes f :: "'a \<Rightarrow> 'b"
+  assumes f: "f \<in> measurable M N" and A: "A \<in> sets N"
+  shows "emeasure (distr M N f) A = emeasure M (f -` A \<inter> space M)" (is "_ = ?\<mu> A")
+  unfolding distr_def
+proof (rule emeasure_measure_of_sigma)
+  show "positive (sets N) ?\<mu>"
+    by (auto simp: positive_def)
+
+  show "countably_additive (sets N) ?\<mu>"
+  proof (intro countably_additiveI)
+    fix A :: "nat \<Rightarrow> 'b set" assume "range A \<subseteq> sets N" "disjoint_family A"
+    then have A: "\<And>i. A i \<in> sets N" "(\<Union>i. A i) \<in> sets N" by auto
+    then have *: "range (\<lambda>i. f -` (A i) \<inter> space M) \<subseteq> sets M"
+      using f by (auto simp: measurable_def)
+    moreover have "(\<Union>i. f -`  A i \<inter> space M) \<in> sets M"
+      using * by blast
+    moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)"
+      using `disjoint_family A` by (auto simp: disjoint_family_on_def)
+    ultimately show "(\<Sum>i. ?\<mu> (A i)) = ?\<mu> (\<Union>i. A i)"
+      using suminf_emeasure[OF _ **] A f
+      by (auto simp: comp_def vimage_UN)
+  qed
+  show "sigma_algebra (space N) (sets N)" ..
+qed fact
+
+lemma AE_distrD:
+  assumes f: "f \<in> measurable M M'"
+    and AE: "AE x in distr M M' f. P x"
+  shows "AE x in M. P (f x)"
+proof -
+  from AE[THEN AE_E] guess N .
+  with f show ?thesis
+    unfolding eventually_ae_filter
+    by (intro bexI[of _ "f -` N \<inter> space M"])
+       (auto simp: emeasure_distr measurable_def)
+qed
+
+lemma null_sets_distr_iff:
+  "f \<in> measurable M N \<Longrightarrow> A \<in> null_sets (distr M N f) \<longleftrightarrow> f -` A \<inter> space M \<in> null_sets M \<and> A \<in> sets N"
+  by (auto simp add: null_sets_def emeasure_distr measurable_sets)
+
+lemma distr_distr:
+  assumes f: "g \<in> measurable N L" and g: "f \<in> measurable M N"
+  shows "distr (distr M N f) L g = distr M L (g \<circ> f)" (is "?L = ?R")
+  using measurable_comp[OF g f] f g
+  by (auto simp add: emeasure_distr measurable_sets measurable_space
+           intro!: arg_cong[where f="emeasure M"] measure_eqI)
+
+section {* Real measure values *}
+
+lemma measure_nonneg: "0 \<le> measure M A"
+  using emeasure_nonneg[of M A] unfolding measure_def by (auto intro: real_of_ereal_pos)
+
+lemma measure_empty[simp]: "measure M {} = 0"
+  unfolding measure_def by simp
+
+lemma emeasure_eq_ereal_measure:
+  "emeasure M A \<noteq> \<infinity> \<Longrightarrow> emeasure M A = ereal (measure M A)"
+  using emeasure_nonneg[of M A]
+  by (cases "emeasure M A") (auto simp: measure_def)
+
+lemma measure_Union:
+  assumes finite: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
+  and measurable: "A \<in> sets M" "B \<in> sets M" "A \<inter> B = {}"
+  shows "measure M (A \<union> B) = measure M A + measure M B"
+  unfolding measure_def
+  using plus_emeasure[OF measurable, symmetric] finite
+  by (simp add: emeasure_eq_ereal_measure)
+
+lemma measure_finite_Union:
+  assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
+  assumes finite: "\<And>i. i \<in> S \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
+  shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
+  unfolding measure_def
+  using setsum_emeasure[OF measurable, symmetric] finite
+  by (simp add: emeasure_eq_ereal_measure)
+
+lemma measure_Diff:
+  assumes finite: "emeasure M A \<noteq> \<infinity>"
+  and measurable: "A \<in> sets M" "B \<in> sets M" "B \<subseteq> A"
+  shows "measure M (A - B) = measure M A - measure M B"
+proof -
+  have "emeasure M (A - B) \<le> emeasure M A" "emeasure M B \<le> emeasure M A"
+    using measurable by (auto intro!: emeasure_mono)
+  hence "measure M ((A - B) \<union> B) = measure M (A - B) + measure M B"
+    using measurable finite by (rule_tac measure_Union) auto
+  thus ?thesis using `B \<subseteq> A` by (auto simp: Un_absorb2)
+qed
+
+lemma measure_UNION:
+  assumes measurable: "range A \<subseteq> sets M" "disjoint_family A"
+  assumes finite: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
+  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
+proof -
+  from summable_sums[OF summable_ereal_pos, of "\<lambda>i. emeasure M (A i)"]
+       suminf_emeasure[OF measurable] emeasure_nonneg[of M]
+  have "(\<lambda>i. emeasure M (A i)) sums (emeasure M (\<Union>i. A i))" by simp
+  moreover
+  { fix i
+    have "emeasure M (A i) \<le> emeasure M (\<Union>i. A i)"
+      using measurable by (auto intro!: emeasure_mono)
+    then have "emeasure M (A i) = ereal ((measure M (A i)))"
+      using finite by (intro emeasure_eq_ereal_measure) auto }
+  ultimately show ?thesis using finite
+    unfolding sums_ereal[symmetric] by (simp add: emeasure_eq_ereal_measure)
+qed
+
+lemma measure_subadditive:
+  assumes measurable: "A \<in> sets M" "B \<in> sets M"
+  and fin: "emeasure M A \<noteq> \<infinity>" "emeasure M B \<noteq> \<infinity>"
+  shows "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
+proof -
+  have "emeasure M (A \<union> B) \<noteq> \<infinity>"
+    using emeasure_subadditive[OF measurable] fin by auto
+  then show "(measure M (A \<union> B)) \<le> (measure M A) + (measure M B)"
+    using emeasure_subadditive[OF measurable] fin
+    by (auto simp: emeasure_eq_ereal_measure)
+qed
+
+lemma measure_subadditive_finite:
+  assumes A: "finite I" "A`I \<subseteq> sets M" and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (A i) \<noteq> \<infinity>"
+  shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
+proof -
+  { have "emeasure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. emeasure M (A i))"
+      using emeasure_subadditive_finite[OF A] .
+    also have "\<dots> < \<infinity>"
+      using fin by (simp add: setsum_Pinfty)
+    finally have "emeasure M (\<Union>i\<in>I. A i) \<noteq> \<infinity>" by simp }
+  then show ?thesis
+    using emeasure_subadditive_finite[OF A] fin
+    unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
+qed
+
+lemma measure_subadditive_countably:
+  assumes A: "range A \<subseteq> sets M" and fin: "(\<Sum>i. emeasure M (A i)) \<noteq> \<infinity>"
+  shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
+proof -
+  from emeasure_nonneg fin have "\<And>i. emeasure M (A i) \<noteq> \<infinity>" by (rule suminf_PInfty)
+  moreover
+  { have "emeasure M (\<Union>i. A i) \<le> (\<Sum>i. emeasure M (A i))"
+      using emeasure_subadditive_countably[OF A] .
+    also have "\<dots> < \<infinity>"
+      using fin by simp
+    finally have "emeasure M (\<Union>i. A i) \<noteq> \<infinity>" by simp }
+  ultimately  show ?thesis
+    using emeasure_subadditive_countably[OF A] fin
+    unfolding measure_def by (simp add: emeasure_eq_ereal_measure suminf_ereal measure_nonneg)
+qed
+
+lemma measure_eq_setsum_singleton:
+  assumes S: "finite S" "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+  and fin: "\<And>x. x \<in> S \<Longrightarrow> emeasure M {x} \<noteq> \<infinity>"
+  shows "(measure M S) = (\<Sum>x\<in>S. (measure M {x}))"
+  unfolding measure_def
+  using emeasure_eq_setsum_singleton[OF S] fin
+  by simp (simp add: emeasure_eq_ereal_measure)
+
+lemma Lim_measure_incseq:
+  assumes A: "range A \<subseteq> sets M" "incseq A" and fin: "emeasure M (\<Union>i. A i) \<noteq> \<infinity>"
+  shows "(\<lambda>i. (measure M (A i))) ----> (measure M (\<Union>i. A i))"
+proof -
+  have "ereal ((measure M (\<Union>i. A i))) = emeasure M (\<Union>i. A i)"
+    using fin by (auto simp: emeasure_eq_ereal_measure)
+  then show ?thesis
+    using Lim_emeasure_incseq[OF A]
+    unfolding measure_def
+    by (intro lim_real_of_ereal) simp
+qed
+
+lemma Lim_measure_decseq:
+  assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
+  shows "(\<lambda>n. measure M (A n)) ----> measure M (\<Inter>i. A i)"
+proof -
+  have "emeasure M (\<Inter>i. A i) \<le> emeasure M (A 0)"
+    using A by (auto intro!: emeasure_mono)
+  also have "\<dots> < \<infinity>"
+    using fin[of 0] by auto
+  finally have "ereal ((measure M (\<Inter>i. A i))) = emeasure M (\<Inter>i. A i)"
+    by (auto simp: emeasure_eq_ereal_measure)
+  then show ?thesis
+    unfolding measure_def
+    using Lim_emeasure_decseq[OF A fin]
+    by (intro lim_real_of_ereal) simp
+qed
+
+section {* Measure spaces with @{term "emeasure M (space M) < \<infinity>"} *}
+
+locale finite_measure = sigma_finite_measure M for M +
+  assumes finite_emeasure_space: "emeasure M (space M) \<noteq> \<infinity>"
+
+lemma finite_measureI[Pure.intro!]:
+  assumes *: "emeasure M (space M) \<noteq> \<infinity>"
+  shows "finite_measure M"
+proof
+  show "\<exists>A. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and> (\<forall>i. emeasure M (A i) \<noteq> \<infinity>)"
+    using * by (auto intro!: exI[of _ "\<lambda>_. space M"])
+qed fact
+
+lemma (in finite_measure) emeasure_finite[simp, intro]: "emeasure M A \<noteq> \<infinity>"
+  using finite_emeasure_space emeasure_space[of M A] by auto
+
+lemma (in finite_measure) emeasure_eq_measure: "emeasure M A = ereal (measure M A)"
+  unfolding measure_def by (simp add: emeasure_eq_ereal_measure)
+
+lemma (in finite_measure) emeasure_real: "\<exists>r. 0 \<le> r \<and> emeasure M A = ereal r"
+  using emeasure_finite[of A] emeasure_nonneg[of M A] by (cases "emeasure M A") auto
+
+lemma (in finite_measure) bounded_measure: "measure M A \<le> measure M (space M)"
+  using emeasure_space[of M A] emeasure_real[of A] emeasure_real[of "space M"] by (auto simp: measure_def)
+
+lemma (in finite_measure) finite_measure_Diff:
+  assumes sets: "A \<in> sets M" "B \<in> sets M" and "B \<subseteq> A"
+  shows "measure M (A - B) = measure M A - measure M B"
+  using measure_Diff[OF _ assms] by simp
+
+lemma (in finite_measure) finite_measure_Union:
+  assumes sets: "A \<in> sets M" "B \<in> sets M" and "A \<inter> B = {}"
+  shows "measure M (A \<union> B) = measure M A + measure M B"
+  using measure_Union[OF _ _ assms] by simp
+
+lemma (in finite_measure) finite_measure_finite_Union:
+  assumes measurable: "A`S \<subseteq> sets M" "disjoint_family_on A S" "finite S"
+  shows "measure M (\<Union>i\<in>S. A i) = (\<Sum>i\<in>S. measure M (A i))"
+  using measure_finite_Union[OF assms] by simp
+
+lemma (in finite_measure) finite_measure_UNION:
+  assumes A: "range A \<subseteq> sets M" "disjoint_family A"
+  shows "(\<lambda>i. measure M (A i)) sums (measure M (\<Union>i. A i))"
+  using measure_UNION[OF A] by simp
+
+lemma (in finite_measure) finite_measure_mono:
+  assumes "A \<subseteq> B" "B \<in> sets M" shows "measure M A \<le> measure M B"
+  using emeasure_mono[OF assms] emeasure_real[of A] emeasure_real[of B] by (auto simp: measure_def)
+
+lemma (in finite_measure) finite_measure_subadditive:
+  assumes m: "A \<in> sets M" "B \<in> sets M"
+  shows "measure M (A \<union> B) \<le> measure M A + measure M B"
+  using measure_subadditive[OF m] by simp
+
+lemma (in finite_measure) finite_measure_subadditive_finite:
+  assumes "finite I" "A`I \<subseteq> sets M" shows "measure M (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. measure M (A i))"
+  using measure_subadditive_finite[OF assms] by simp
+
+lemma (in finite_measure) finite_measure_subadditive_countably:
+  assumes A: "range A \<subseteq> sets M" and sum: "summable (\<lambda>i. measure M (A i))"
+  shows "measure M (\<Union>i. A i) \<le> (\<Sum>i. measure M (A i))"
+proof -
+  from `summable (\<lambda>i. measure M (A i))`
+  have "(\<lambda>i. ereal (measure M (A i))) sums ereal (\<Sum>i. measure M (A i))"
+    by (simp add: sums_ereal) (rule summable_sums)
+  from sums_unique[OF this, symmetric]
+       measure_subadditive_countably[OF A]
+  show ?thesis by (simp add: emeasure_eq_measure)
+qed
+
+lemma (in finite_measure) finite_measure_eq_setsum_singleton:
+  assumes "finite S" and *: "\<And>x. x \<in> S \<Longrightarrow> {x} \<in> sets M"
+  shows "measure M S = (\<Sum>x\<in>S. measure M {x})"
+  using measure_eq_setsum_singleton[OF assms] by simp
+
+lemma (in finite_measure) finite_Lim_measure_incseq:
+  assumes A: "range A \<subseteq> sets M" "incseq A"
+  shows "(\<lambda>i. measure M (A i)) ----> measure M (\<Union>i. A i)"
+  using Lim_measure_incseq[OF A] by simp
+
+lemma (in finite_measure) finite_Lim_measure_decseq:
+  assumes A: "range A \<subseteq> sets M" "decseq A"
+  shows "(\<lambda>n. measure M (A n)) ----> measure M (\<Inter>i. A i)"
+  using Lim_measure_decseq[OF A] by simp
+
+lemma (in finite_measure) finite_measure_compl:
+  assumes S: "S \<in> sets M"
+  shows "measure M (space M - S) = measure M (space M) - measure M S"
+  using measure_Diff[OF _ top S sets_into_space] S by simp
+
+lemma (in finite_measure) finite_measure_mono_AE:
+  assumes imp: "AE x in M. x \<in> A \<longrightarrow> x \<in> B" and B: "B \<in> sets M"
+  shows "measure M A \<le> measure M B"
+  using assms emeasure_mono_AE[OF imp B]
+  by (simp add: emeasure_eq_measure)
+
+lemma (in finite_measure) finite_measure_eq_AE:
+  assumes iff: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
+  assumes A: "A \<in> sets M" and B: "B \<in> sets M"
+  shows "measure M A = measure M B"
+  using assms emeasure_eq_AE[OF assms] by (simp add: emeasure_eq_measure)
+
+section {* Counting space *}
+
+definition count_space :: "'a set \<Rightarrow> 'a measure" where
+  "count_space \<Omega> = measure_of \<Omega> (Pow \<Omega>) (\<lambda>A. if finite A then ereal (card A) else \<infinity>)"
+
+lemma 
+  shows space_count_space[simp]: "space (count_space \<Omega>) = \<Omega>"
+    and sets_count_space[simp]: "sets (count_space \<Omega>) = Pow \<Omega>"
+  using sigma_sets_into_sp[of "Pow \<Omega>" \<Omega>]
+  by (auto simp: count_space_def)
+
+lemma measurable_count_space_eq1[simp]:
+  "f \<in> measurable (count_space A) M \<longleftrightarrow> f \<in> A \<rightarrow> space M"
+ unfolding measurable_def by simp
+
+lemma measurable_count_space_eq2[simp]:
+  assumes "finite A"
+  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
+proof -
+  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
+    with `finite A` have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "finite X"
+      by (auto dest: finite_subset)
+    moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
+    ultimately have "f -` X \<inter> space M \<in> sets M"
+      using `X \<subseteq> A` by (auto intro!: finite_UN simp del: UN_simps) }
+  then show ?thesis
+    unfolding measurable_def by auto
+qed
+
+lemma emeasure_count_space:
+  assumes "X \<subseteq> A" shows "emeasure (count_space A) X = (if finite X then ereal (card X) else \<infinity>)"
+    (is "_ = ?M X")
+  unfolding count_space_def
+proof (rule emeasure_measure_of_sigma)
+  show "sigma_algebra A (Pow A)" by (rule sigma_algebra_Pow)
+
+  show "positive (Pow A) ?M"
+    by (auto simp: positive_def)
+
+  show "countably_additive (Pow A) ?M"
+  proof (unfold countably_additive_def, safe)
+      fix F :: "nat \<Rightarrow> 'a set" assume disj: "disjoint_family F"
+      show "(\<Sum>i. ?M (F i)) = ?M (\<Union>i. F i)"
+      proof cases
+        assume "\<forall>i. finite (F i)"
+        then have finite_F: "\<And>i. finite (F i)" by auto
+        have "\<forall>i\<in>{i. F i \<noteq> {}}. \<exists>x. x \<in> F i" by auto
+        from bchoice[OF this] obtain f where f: "\<And>i. F i \<noteq> {} \<Longrightarrow> f i \<in> F i" by auto
+
+        have inj_f: "inj_on f {i. F i \<noteq> {}}"
+        proof (rule inj_onI, simp)
+          fix i j a b assume *: "f i = f j" "F i \<noteq> {}" "F j \<noteq> {}"
+          then have "f i \<in> F i" "f j \<in> F j" using f by force+
+          with disj * show "i = j" by (auto simp: disjoint_family_on_def)
+        qed
+        have fin_eq: "finite (\<Union>i. F i) \<longleftrightarrow> finite {i. F i \<noteq> {}}"
+        proof
+          assume "finite (\<Union>i. F i)"
+          show "finite {i. F i \<noteq> {}}"
+          proof (rule finite_imageD)
+            from f have "f`{i. F i \<noteq> {}} \<subseteq> (\<Union>i. F i)" by auto
+            then show "finite (f`{i. F i \<noteq> {}})"
+              by (rule finite_subset) fact
+          qed fact
+        next
+          assume "finite {i. F i \<noteq> {}}"
+          with finite_F have "finite (\<Union>i\<in>{i. F i \<noteq> {}}. F i)"
+            by auto
+          also have "(\<Union>i\<in>{i. F i \<noteq> {}}. F i) = (\<Union>i. F i)"
+            by auto
+          finally show "finite (\<Union>i. F i)" .
+        qed
+        
+        show ?thesis
+        proof cases
+          assume *: "finite (\<Union>i. F i)"
+          with finite_F have "finite {i. ?M (F i) \<noteq> 0} "
+            by (simp add: fin_eq)
+          then have "(\<Sum>i. ?M (F i)) = (\<Sum>i | ?M (F i) \<noteq> 0. ?M (F i))"
+            by (rule suminf_eq_setsum)
+          also have "\<dots> = ereal (\<Sum>i | F i \<noteq> {}. card (F i))"
+            using finite_F by simp
+          also have "\<dots> = ereal (card (\<Union>i \<in> {i. F i \<noteq> {}}. F i))"
+            using * finite_F disj
+            by (subst card_UN_disjoint) (auto simp: disjoint_family_on_def fin_eq)
+          also have "\<dots> = ?M (\<Union>i. F i)"
+            using * by (auto intro!: arg_cong[where f=card])
+          finally show ?thesis .
+        next
+          assume inf: "infinite (\<Union>i. F i)"
+          { fix i
+            have "\<exists>N. i \<le> (\<Sum>i<N. card (F i))"
+            proof (induct i)
+              case (Suc j)
+              from Suc obtain N where N: "j \<le> (\<Sum>i<N. card (F i))" by auto
+              have "infinite ({i. F i \<noteq> {}} - {..< N})"
+                using inf by (auto simp: fin_eq)
+              then have "{i. F i \<noteq> {}} - {..< N} \<noteq> {}"
+                by (metis finite.emptyI)
+              then obtain i where i: "F i \<noteq> {}" "N \<le> i"
+                by (auto simp: not_less[symmetric])
+
+              note N
+              also have "(\<Sum>i<N. card (F i)) \<le> (\<Sum>i<i. card (F i))"
+                by (rule setsum_mono2) (auto simp: i)
+              also have "\<dots> < (\<Sum>i<i. card (F i)) + card (F i)"
+                using finite_F `F i \<noteq> {}` by (simp add: card_gt_0_iff)
+              finally have "j < (\<Sum>i<Suc i. card (F i))"
+                by simp
+              then show ?case unfolding Suc_le_eq by blast
+            qed simp }
+          with finite_F inf show ?thesis
+            by (auto simp del: real_of_nat_setsum intro!: SUP_PInfty
+                     simp add: suminf_ereal_eq_SUPR real_of_nat_setsum[symmetric])
+        qed
+      next
+        assume "\<not> (\<forall>i. finite (F i))"
+        then obtain j where j: "infinite (F j)" by auto
+        then have "infinite (\<Union>i. F i)"
+          using finite_subset[of "F j" "\<Union>i. F i"] by auto
+        moreover have "\<And>i. 0 \<le> ?M (F i)" by auto
+        ultimately show ?thesis
+          using suminf_PInfty[of "\<lambda>i. ?M (F i)" j] j by auto
+      qed
+  qed
+  show "X \<in> Pow A" using `X \<subseteq> A` by simp
+qed
+
+lemma emeasure_count_space_finite[simp]:
+  "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> emeasure (count_space A) X = ereal (card X)"
+  using emeasure_count_space[of X A] by simp
+
+lemma emeasure_count_space_infinite[simp]:
+  "X \<subseteq> A \<Longrightarrow> infinite X \<Longrightarrow> emeasure (count_space A) X = \<infinity>"
+  using emeasure_count_space[of X A] by simp
+
+lemma emeasure_count_space_eq_0:
+  "emeasure (count_space A) X = 0 \<longleftrightarrow> (X \<subseteq> A \<longrightarrow> X = {})"
+proof cases
+  assume X: "X \<subseteq> A"
+  then show ?thesis
+  proof (intro iffI impI)
+    assume "emeasure (count_space A) X = 0"
+    with X show "X = {}"
+      by (subst (asm) emeasure_count_space) (auto split: split_if_asm)
+  qed simp
+qed (simp add: emeasure_notin_sets)
+
+lemma null_sets_count_space: "null_sets (count_space A) = { {} }"
+  unfolding null_sets_def by (auto simp add: emeasure_count_space_eq_0)
+
+lemma AE_count_space: "(AE x in count_space A. P x) \<longleftrightarrow> (\<forall>x\<in>A. P x)"
+  unfolding eventually_ae_filter by (auto simp add: null_sets_count_space)
+
+lemma sigma_finite_measure_count_space:
+  fixes A :: "'a::countable set"
+  shows "sigma_finite_measure (count_space A)"
+proof
+  show "\<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (count_space A) \<and> (\<Union>i. F i) = space (count_space A) \<and>
+     (\<forall>i. emeasure (count_space A) (F i) \<noteq> \<infinity>)"
+     using surj_from_nat by (intro exI[of _ "\<lambda>i. {from_nat i} \<inter> A"]) (auto simp del: surj_from_nat)
+qed
+
+lemma finite_measure_count_space:
+  assumes [simp]: "finite A"
+  shows "finite_measure (count_space A)"
+  by rule simp
+
+lemma sigma_finite_measure_count_space_finite:
+  assumes A: "finite A" shows "sigma_finite_measure (count_space A)"
+proof -
+  interpret finite_measure "count_space A" using A by (rule finite_measure_count_space)
+  show "sigma_finite_measure (count_space A)" ..
+qed
+
+end
+
--- a/src/HOL/Probability/Probability.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Probability.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -4,7 +4,7 @@
   Probability_Measure
   Infinite_Product_Measure
   Independent_Family
-  Conditional_Probability
   Information
 begin
+
 end
--- a/src/HOL/Probability/Probability_Measure.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Probability_Measure.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -6,110 +6,219 @@
 header {*Probability measure*}
 
 theory Probability_Measure
-imports Lebesgue_Measure
+  imports Lebesgue_Measure Radon_Nikodym
 begin
 
+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_funcset_insert_eq[simp]:
+  assumes "a \<notin> A"
+  shows "extensional (insert a A) \<inter> (insert a A \<rightarrow> B) = (\<Union>f \<in> extensional 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 A \<inter> (A \<rightarrow> B))"
+  using assms by induct auto
+
+lemma finite_PiE[simp, intro]:
+  assumes fin: "finite A" "\<And>i. i \<in> A \<Longrightarrow> finite (B i)"
+  shows "finite (Pi\<^isub>E A B)"
+proof -
+  have *: "(Pi\<^isub>E A B) \<subseteq> extensional A \<inter> (A \<rightarrow> (\<Union>i\<in>A. B i))" by auto
+  show ?thesis
+    using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto
+qed
+
+
+lemma countably_additiveI[case_names countably]:
+  assumes "\<And>A. \<lbrakk> range A \<subseteq> M ; disjoint_family A ; (\<Union>i. A i) \<in> M\<rbrakk> \<Longrightarrow> (\<Sum>n. \<mu> (A n)) = \<mu> (\<Union>i. A i)"
+  shows "countably_additive M \<mu>"
+  using assms unfolding countably_additive_def by auto
+
+lemma convex_le_Inf_differential:
+  fixes f :: "real \<Rightarrow> real"
+  assumes "convex_on I f"
+  assumes "x \<in> interior I" "y \<in> I"
+  shows "f y \<ge> f x + Inf ((\<lambda>t. (f x - f t) / (x - t)) ` ({x<..} \<inter> I)) * (y - x)"
+    (is "_ \<ge> _ + Inf (?F x) * (y - x)")
+proof -
+  show ?thesis
+  proof (cases rule: linorder_cases)
+    assume "x < y"
+    moreover
+    have "open (interior I)" by auto
+    from openE[OF this `x \<in> interior I`] guess e . note e = this
+    moreover def t \<equiv> "min (x + e / 2) ((x + y) / 2)"
+    ultimately have "x < t" "t < y" "t \<in> ball x e"
+      by (auto simp: dist_real_def field_simps split: split_min)
+    with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
+
+    have "open (interior I)" by auto
+    from openE[OF this `x \<in> interior I`] guess e .
+    moreover def K \<equiv> "x - e / 2"
+    with `0 < e` have "K \<in> ball x e" "K < x" by (auto simp: dist_real_def)
+    ultimately have "K \<in> I" "K < x" "x \<in> I"
+      using interior_subset[of I] `x \<in> interior I` by auto
+
+    have "Inf (?F x) \<le> (f x - f y) / (x - y)"
+    proof (rule Inf_lower2)
+      show "(f x - f t) / (x - t) \<in> ?F x"
+        using `t \<in> I` `x < t` by auto
+      show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
+        using `convex_on I f` `x \<in> I` `y \<in> I` `x < t` `t < y` by (rule convex_on_diff)
+    next
+      fix y assume "y \<in> ?F x"
+      with order_trans[OF convex_on_diff[OF `convex_on I f` `K \<in> I` _ `K < x` _]]
+      show "(f K - f x) / (K - x) \<le> y" by auto
+    qed
+    then show ?thesis
+      using `x < y` by (simp add: field_simps)
+  next
+    assume "y < x"
+    moreover
+    have "open (interior I)" by auto
+    from openE[OF this `x \<in> interior I`] guess e . note e = this
+    moreover def t \<equiv> "x + e / 2"
+    ultimately have "x < t" "t \<in> ball x e"
+      by (auto simp: dist_real_def field_simps)
+    with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
+
+    have "(f x - f y) / (x - y) \<le> Inf (?F x)"
+    proof (rule Inf_greatest)
+      have "(f x - f y) / (x - y) = (f y - f x) / (y - x)"
+        using `y < x` by (auto simp: field_simps)
+      also
+      fix z  assume "z \<in> ?F x"
+      with order_trans[OF convex_on_diff[OF `convex_on I f` `y \<in> I` _ `y < x`]]
+      have "(f y - f x) / (y - x) \<le> z" by auto
+      finally show "(f x - f y) / (x - y) \<le> z" .
+    next
+      have "open (interior I)" by auto
+      from openE[OF this `x \<in> interior I`] guess e . note e = this
+      then have "x + e / 2 \<in> ball x e" by (auto simp: dist_real_def)
+      with e interior_subset[of I] have "x + e / 2 \<in> {x<..} \<inter> I" by auto
+      then show "?F x \<noteq> {}" by blast
+    qed
+    then show ?thesis
+      using `y < x` by (simp add: field_simps)
+  qed simp
+qed
+
+lemma distr_id[simp]: "distr N N (\<lambda>x. x) = N"
+  by (rule measure_eqI) (auto simp: emeasure_distr)
+
 locale prob_space = finite_measure +
-  assumes measure_space_1: "measure M (space M) = 1"
+  assumes emeasure_space_1: "emeasure M (space M) = 1"
 
 lemma prob_spaceI[Pure.intro!]:
-  assumes "measure_space M"
-  assumes *: "measure M (space M) = 1"
+  assumes *: "emeasure M (space M) = 1"
   shows "prob_space M"
 proof -
   interpret finite_measure M
   proof
-    show "measure_space M" by fact
-    show "measure M (space M) \<noteq> \<infinity>" using * by simp 
+    show "emeasure M (space M) \<noteq> \<infinity>" using * by simp 
   qed
   show "prob_space M" by default fact
 qed
 
 abbreviation (in prob_space) "events \<equiv> sets M"
-abbreviation (in prob_space) "prob \<equiv> \<mu>'"
-abbreviation (in prob_space) "random_variable M' X \<equiv> sigma_algebra M' \<and> X \<in> measurable M M'"
+abbreviation (in prob_space) "prob \<equiv> measure M"
+abbreviation (in prob_space) "random_variable M' X \<equiv> X \<in> measurable M M'"
 abbreviation (in prob_space) "expectation \<equiv> integral\<^isup>L M"
 
-definition (in prob_space)
-  "distribution X A = \<mu>' (X -` A \<inter> space M)"
-
-abbreviation (in prob_space)
-  "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
-
-lemma (in prob_space) prob_space_cong:
-  assumes "\<And>A. A \<in> sets M \<Longrightarrow> measure N A = \<mu> A" "space N = space M" "sets N = sets M"
-  shows "prob_space N"
-proof
-  show "measure_space N" by (intro measure_space_cong assms)
-  show "measure N (space N) = 1"
-    using measure_space_1 assms by simp
+lemma (in prob_space) prob_space_distr:
+  assumes f: "f \<in> measurable M M'" shows "prob_space (distr M M' f)"
+proof (rule prob_spaceI)
+  have "f -` space M' \<inter> space M = space M" using f by (auto dest: measurable_space)
+  with f show "emeasure (distr M M' f) (space (distr M M' f)) = 1"
+    by (auto simp: emeasure_distr emeasure_space_1)
 qed
 
-lemma (in prob_space) distribution_cong:
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
-  shows "distribution X = distribution Y"
-  unfolding distribution_def fun_eq_iff
-  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
-
-lemma (in prob_space) joint_distribution_cong:
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
-  shows "joint_distribution X Y = joint_distribution X' Y'"
-  unfolding distribution_def fun_eq_iff
-  using assms by (auto intro!: arg_cong[where f="\<mu>'"])
-
-lemma (in prob_space) distribution_id[simp]:
-  "N \<in> events \<Longrightarrow> distribution (\<lambda>x. x) N = prob N"
-  by (auto simp: distribution_def intro!: arg_cong[where f=prob])
-
 lemma (in prob_space) prob_space: "prob (space M) = 1"
-  using measure_space_1 unfolding \<mu>'_def by (simp add: one_ereal_def)
+  using emeasure_space_1 unfolding measure_def by (simp add: one_ereal_def)
 
 lemma (in prob_space) prob_le_1[simp, intro]: "prob A \<le> 1"
   using bounded_measure[of A] by (simp add: prob_space)
 
-lemma (in prob_space) distribution_positive[simp, intro]:
-  "0 \<le> distribution X A" unfolding distribution_def by auto
-
-lemma (in prob_space) not_zero_less_distribution[simp]:
-  "(\<not> 0 < distribution X A) \<longleftrightarrow> distribution X A = 0"
-  using distribution_positive[of X A] by arith
-
-lemma (in prob_space) joint_distribution_remove[simp]:
-    "joint_distribution X X {(x, x)} = distribution X {x}"
-  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>'])
+lemma (in prob_space) not_empty: "space M \<noteq> {}"
+  using prob_space by auto
 
-lemma (in prob_space) distribution_unit[simp]: "distribution (\<lambda>x. ()) {()} = 1"
-  unfolding distribution_def using prob_space 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) not_empty: "space M \<noteq> {}"
-  using prob_space empty_measure' by auto
-
-lemma (in prob_space) measure_le_1: "X \<in> sets M \<Longrightarrow> \<mu> X \<le> 1"
-  unfolding measure_space_1[symmetric]
-  using sets_into_space
-  by (intro measure_mono) auto
+lemma (in prob_space) measure_le_1: "emeasure M X \<le> 1"
+  using emeasure_space[of M X] by (simp add: emeasure_space_1)
 
 lemma (in prob_space) AE_I_eq_1:
-  assumes "\<mu> {x\<in>space M. P x} = 1" "{x\<in>space M. P x} \<in> sets M"
-  shows "AE x. P x"
+  assumes "emeasure M {x\<in>space M. P x} = 1" "{x\<in>space M. P x} \<in> sets M"
+  shows "AE x in M. P x"
 proof (rule AE_I)
-  show "\<mu> (space M - {x \<in> space M. P x}) = 0"
-    using assms measure_space_1 by (simp add: measure_compl)
+  show "emeasure M (space M - {x \<in> space M. P x}) = 0"
+    using assms emeasure_space_1 by (simp add: emeasure_compl)
 qed (insert assms, auto)
 
-lemma (in prob_space) distribution_1:
-  "distribution X A \<le> 1"
-  unfolding distribution_def by simp
-
 lemma (in prob_space) prob_compl:
   assumes A: "A \<in> events"
   shows "prob (space M - A) = 1 - prob A"
   using finite_measure_compl[OF A] by (simp add: prob_space)
 
+lemma (in prob_space) AE_in_set_eq_1:
+  assumes "A \<in> events" shows "(AE x in M. x \<in> A) \<longleftrightarrow> prob A = 1"
+proof
+  assume ae: "AE x in M. x \<in> A"
+  have "{x \<in> space M. x \<in> A} = A" "{x \<in> space M. x \<notin> A} = space M - A"
+    using `A \<in> events`[THEN sets_into_space] by auto
+  with AE_E2[OF ae] `A \<in> events` have "1 - emeasure M A = 0"
+    by (simp add: emeasure_compl emeasure_space_1)
+  then show "prob A = 1"
+    using `A \<in> events` by (simp add: emeasure_eq_measure one_ereal_def)
+next
+  assume prob: "prob A = 1"
+  show "AE x in M. x \<in> A"
+  proof (rule AE_I)
+    show "{x \<in> space M. x \<notin> A} \<subseteq> space M - A" by auto
+    show "emeasure M (space M - A) = 0"
+      using `A \<in> events` prob
+      by (simp add: prob_compl emeasure_space_1 emeasure_eq_measure one_ereal_def)
+    show "space M - A \<in> events"
+      using `A \<in> events` by auto
+  qed
+qed
+
+lemma (in prob_space) AE_False: "(AE x in M. False) \<longleftrightarrow> False"
+proof
+  assume "AE x in M. False"
+  then have "AE x in M. x \<in> {}" by simp
+  then show False
+    by (subst (asm) AE_in_set_eq_1) auto
+qed simp
+
+lemma (in prob_space) AE_prob_1:
+  assumes "prob A = 1" shows "AE x in M. x \<in> A"
+proof -
+  from `prob A = 1` have "A \<in> events"
+    by (metis measure_notin_sets zero_neq_one)
+  with AE_in_set_eq_1 assms show ?thesis by simp
+qed
+
 lemma (in prob_space) prob_space_increasing: "increasing M prob"
   by (auto intro!: finite_measure_mono simp: increasing_def)
 
@@ -164,9 +273,8 @@
   shows "prob (\<Union> i :: nat. c i) = 0"
 proof (rule antisym)
   show "prob (\<Union> i :: nat. c i) \<le> 0"
-    using finite_measure_countably_subadditive[OF assms(1)]
-    by (simp add: assms(2) suminf_zero summable_zero)
-qed simp
+    using finite_measure_subadditive_countably[OF assms(1)] by (simp add: assms(2))
+qed (simp add: measure_nonneg)
 
 lemma (in prob_space) prob_equiprobable_finite_unions:
   assumes "s \<in> events"
@@ -178,7 +286,7 @@
   from someI_ex[OF this] assms
   have prob_some: "\<And> x. x \<in> s \<Longrightarrow> prob {x} = prob {SOME y. y \<in> s}" by blast
   have "prob s = (\<Sum> x \<in> s. prob {x})"
-    using finite_measure_finite_singleton[OF s_finite] by simp
+    using finite_measure_eq_setsum_singleton[OF s_finite] by simp
   also have "\<dots> = (\<Sum> x \<in> s. prob {SOME y. y \<in> s})" using prob_some by auto
   also have "\<dots> = real (card s) * prob {(SOME x. x \<in> s)}"
     using setsum_constant assms by (simp add: real_eq_of_nat)
@@ -199,96 +307,20 @@
   also have "\<dots> = (\<Sum> x \<in> s. prob (e \<inter> f x))"
   proof (rule finite_measure_finite_Union)
     show "finite s" by fact
-    show "\<And>i. i \<in> s \<Longrightarrow> e \<inter> f i \<in> events" by fact
+    show "(\<lambda>i. e \<inter> f i)`s \<subseteq> events" using assms(2) by auto
     show "disjoint_family_on (\<lambda>i. e \<inter> f i) s"
       using disjoint by (auto simp: disjoint_family_on_def)
   qed
   finally show ?thesis .
 qed
 
-lemma (in prob_space) prob_space_vimage:
-  assumes S: "sigma_algebra S"
-  assumes T: "T \<in> measure_preserving M S"
-  shows "prob_space S"
-proof
-  interpret S: measure_space S
-    using S and T by (rule measure_space_vimage)
-  show "measure_space S" ..
-  
-  from T[THEN measure_preservingD2]
-  have "T -` space S \<inter> space M = space M"
-    by (auto simp: measurable_def)
-  with T[THEN measure_preservingD, of "space S", symmetric]
-  show  "measure S (space S) = 1"
-    using measure_space_1 by simp
-qed
-
-lemma prob_space_unique_Int_stable:
-  fixes E :: "('a, 'b) algebra_scheme" and A :: "nat \<Rightarrow> 'a set"
-  assumes E: "Int_stable E" "space E \<in> sets E"
-  and M: "prob_space M" "space M = space E" "sets M = sets (sigma E)"
-  and N: "prob_space N" "space N = space E" "sets N = sets (sigma E)"
-  and eq: "\<And>X. X \<in> sets E \<Longrightarrow> finite_measure.\<mu>' M X = finite_measure.\<mu>' N X"
-  assumes "X \<in> sets (sigma E)"
-  shows "finite_measure.\<mu>' M X = finite_measure.\<mu>' N X"
-proof -
-  interpret M!: prob_space M by fact
-  interpret N!: prob_space N by fact
-  have "measure M X = measure N X"
-  proof (rule measure_unique_Int_stable[OF `Int_stable E`])
-    show "range (\<lambda>i. space M) \<subseteq> sets E" "incseq (\<lambda>i. space M)" "(\<Union>i. space M) = space E"
-      using E M N by auto
-    show "\<And>i. M.\<mu> (space M) \<noteq> \<infinity>"
-      using M.measure_space_1 by simp
-    show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure_space.measure = M.\<mu>\<rparr>"
-      using E M N by (auto intro!: M.measure_space_cong)
-    show "measure_space \<lparr>space = space E, sets = sets (sigma E), measure_space.measure = N.\<mu>\<rparr>"
-      using E M N by (auto intro!: N.measure_space_cong)
-    { fix X assume "X \<in> sets E"
-      then have "X \<in> sets (sigma E)"
-        by (auto simp: sets_sigma sigma_sets.Basic)
-      with eq[OF `X \<in> sets E`] M N show "M.\<mu> X = N.\<mu> X"
-        by (simp add: M.finite_measure_eq N.finite_measure_eq) }
-  qed fact
-  with `X \<in> sets (sigma E)` M N show ?thesis
-    by (simp add: M.finite_measure_eq N.finite_measure_eq)
-qed
-
-lemma (in prob_space) distribution_prob_space:
-  assumes X: "random_variable S X"
-  shows "prob_space (S\<lparr>measure := ereal \<circ> distribution X\<rparr>)" (is "prob_space ?S")
-proof (rule prob_space_vimage)
-  show "X \<in> measure_preserving M ?S"
-    using X
-    unfolding measure_preserving_def distribution_def [abs_def]
-    by (auto simp: finite_measure_eq measurable_sets)
-  show "sigma_algebra ?S" using X by simp
-qed
-
-lemma (in prob_space) AE_distribution:
-  assumes X: "random_variable MX X" and "AE x in MX\<lparr>measure := ereal \<circ> distribution X\<rparr>. Q x"
-  shows "AE x. Q (X x)"
-proof -
-  interpret X: prob_space "MX\<lparr>measure := ereal \<circ> distribution X\<rparr>" 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
-  from X[unfolded measurable_def] N show "AE x. Q (X x)"
-    by (intro AE_I'[where N="X -` N \<inter> space M"])
-       (auto simp: finite_measure_eq distribution_def measurable_sets)
-qed
-
-lemma (in prob_space) distribution_eq_integral:
-  "random_variable S X \<Longrightarrow> A \<in> sets S \<Longrightarrow> distribution X A = expectation (indicator (X -` A \<inter> space M))"
-  using finite_measure_eq[of "X -` A \<inter> space M"]
-  by (auto simp: measurable_sets distribution_def)
-
 lemma (in prob_space) expectation_less:
   assumes [simp]: "integrable M X"
   assumes gt: "\<forall>x\<in>space M. X x < b"
   shows "expectation X < b"
 proof -
   have "expectation X < expectation (\<lambda>x. b)"
-    using gt measure_space_1
+    using gt emeasure_space_1
     by (intro integral_less_AE_space) auto
   then show ?thesis using prob_space by simp
 qed
@@ -299,80 +331,11 @@
   shows "a < expectation X"
 proof -
   have "expectation (\<lambda>x. a) < expectation X"
-    using gt measure_space_1
+    using gt emeasure_space_1
     by (intro integral_less_AE_space) auto
   then show ?thesis using prob_space by simp
 qed
 
-lemma convex_le_Inf_differential:
-  fixes f :: "real \<Rightarrow> real"
-  assumes "convex_on I f"
-  assumes "x \<in> interior I" "y \<in> I"
-  shows "f y \<ge> f x + Inf ((\<lambda>t. (f x - f t) / (x - t)) ` ({x<..} \<inter> I)) * (y - x)"
-    (is "_ \<ge> _ + Inf (?F x) * (y - x)")
-proof -
-  show ?thesis
-  proof (cases rule: linorder_cases)
-    assume "x < y"
-    moreover
-    have "open (interior I)" by auto
-    from openE[OF this `x \<in> interior I`] guess e . note e = this
-    moreover def t \<equiv> "min (x + e / 2) ((x + y) / 2)"
-    ultimately have "x < t" "t < y" "t \<in> ball x e"
-      by (auto simp: mem_ball dist_real_def field_simps split: split_min)
-    with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
-
-    have "open (interior I)" by auto
-    from openE[OF this `x \<in> interior I`] guess e .
-    moreover def K \<equiv> "x - e / 2"
-    with `0 < e` have "K \<in> ball x e" "K < x" by (auto simp: mem_ball dist_real_def)
-    ultimately have "K \<in> I" "K < x" "x \<in> I"
-      using interior_subset[of I] `x \<in> interior I` by auto
-
-    have "Inf (?F x) \<le> (f x - f y) / (x - y)"
-    proof (rule Inf_lower2)
-      show "(f x - f t) / (x - t) \<in> ?F x"
-        using `t \<in> I` `x < t` by auto
-      show "(f x - f t) / (x - t) \<le> (f x - f y) / (x - y)"
-        using `convex_on I f` `x \<in> I` `y \<in> I` `x < t` `t < y` by (rule convex_on_diff)
-    next
-      fix y assume "y \<in> ?F x"
-      with order_trans[OF convex_on_diff[OF `convex_on I f` `K \<in> I` _ `K < x` _]]
-      show "(f K - f x) / (K - x) \<le> y" by auto
-    qed
-    then show ?thesis
-      using `x < y` by (simp add: field_simps)
-  next
-    assume "y < x"
-    moreover
-    have "open (interior I)" by auto
-    from openE[OF this `x \<in> interior I`] guess e . note e = this
-    moreover def t \<equiv> "x + e / 2"
-    ultimately have "x < t" "t \<in> ball x e"
-      by (auto simp: mem_ball dist_real_def field_simps)
-    with `x \<in> interior I` e interior_subset[of I] have "t \<in> I" "x \<in> I" by auto
-
-    have "(f x - f y) / (x - y) \<le> Inf (?F x)"
-    proof (rule Inf_greatest)
-      have "(f x - f y) / (x - y) = (f y - f x) / (y - x)"
-        using `y < x` by (auto simp: field_simps)
-      also
-      fix z  assume "z \<in> ?F x"
-      with order_trans[OF convex_on_diff[OF `convex_on I f` `y \<in> I` _ `y < x`]]
-      have "(f y - f x) / (y - x) \<le> z" by auto
-      finally show "(f x - f y) / (x - y) \<le> z" .
-    next
-      have "open (interior I)" by auto
-      from openE[OF this `x \<in> interior I`] guess e . note e = this
-      then have "x + e / 2 \<in> ball x e" by (auto simp: mem_ball dist_real_def)
-      with e interior_subset[of I] have "x + e / 2 \<in> {x<..} \<inter> I" by auto
-      then show "?F x \<noteq> {}" by blast
-    qed
-    then show ?thesis
-      using `y < x` by (simp add: field_simps)
-  qed simp
-qed
-
 lemma (in prob_space) jensens_inequality:
   fixes a b :: real
   assumes X: "integrable M X" "X ` space M \<subseteq> I"
@@ -410,8 +373,7 @@
     fix k assume "k \<in> (\<lambda>x. q x + ?F x * (expectation X - x)) ` I"
     then guess x .. note x = this
     have "q x + ?F x * (expectation X  - x) = expectation (\<lambda>w. q x + ?F x * (X w - x))"
-      using prob_space
-      by (simp add: integral_add integral_cmult integral_diff lebesgue_integral_const X)
+      using prob_space by (simp add: X)
     also have "\<dots> \<le> expectation (\<lambda>w. q (X w))"
       using `x \<in> I` `open I` X(2)
       by (intro integral_mono integral_add integral_cmult integral_diff
@@ -422,31 +384,6 @@
   finally show "q (expectation X) \<le> expectation (\<lambda>x. q (X x))" .
 qed
 
-lemma (in prob_space) distribution_eq_translated_integral:
-  assumes "random_variable S X" "A \<in> sets S"
-  shows "distribution X A = integral\<^isup>P (S\<lparr>measure := ereal \<circ> distribution X\<rparr>) (indicator A)"
-proof -
-  interpret S: prob_space "S\<lparr>measure := ereal \<circ> distribution X\<rparr>"
-    using assms(1) by (rule distribution_prob_space)
-  show ?thesis
-    using S.positive_integral_indicator(1)[of A] assms by simp
-qed
-
-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))" (is "_ = ?r")
-proof (subst integral_on_finite)
-  show "X \<in> borel_measurable M" "finite (X`space M)" using assms by auto
-  show "(\<Sum> r \<in> X ` space M. r * real (\<mu> (X -` {r} \<inter> space M))) = ?r"
-    "\<And>x. \<mu> (X -` {x} \<inter> space M) \<noteq> \<infinity>"
-    using finite_measure_eq[OF borel_measurable_vimage, of X] rv by auto
-qed
-
-lemma (in prob_space) finite_expectation:
-  assumes "finite (X`space M)" "random_variable borel X"
-  shows "expectation X = (\<Sum> r \<in> X ` (space M). r * distribution X {r})"
-  using assms unfolding distribution_def using finite_expectation1 by auto
-
 lemma (in prob_space) prob_x_eq_1_imp_prob_y_eq_0:
   assumes "{x} \<in> events"
   assumes "prob {x} = 1"
@@ -455,119 +392,25 @@
   shows "prob {y} = 0"
   using prob_one_inter[of "{y}" "{x}"] assms by auto
 
-lemma (in prob_space) distribution_empty[simp]: "distribution X {} = 0"
-  unfolding distribution_def by simp
-
-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: prob_space)
-qed
-
-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
-    using assms[unfolded measurable_def] by (auto intro!: finite_measure_mono)
-  thus ?thesis by (simp add: prob_space)
-qed
-
-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 cases
-  { fix x have "X -` {x} \<inter> space M \<in> sets M"
-    proof cases
-      assume "x \<in> X`space M" with X show ?thesis
-        by (auto simp: measurable_def image_iff)
-    next
-      assume "x \<notin> X`space M" then have "X -` {x} \<inter> space M = {}" by auto
-      then show ?thesis by auto
-    qed } note single = this
-  have "X -` {x} \<inter> space M - X -` {y} \<inter> space M = X -` {x} \<inter> space M"
-    "X -` {y} \<inter> space M \<inter> (X -` {x} \<inter> space M) = {}"
-    using `y \<noteq> x` by auto
-  with finite_measure_inter_full_set[OF single single, of x y] assms(2)
-  show ?thesis by (auto simp: distribution_def prob_space)
-next
-  assume "{y} \<notin> sets ?S"
-  then have "X -` {y} \<inter> space M = {}" by auto
-  thus "distribution X {y} = 0" unfolding distribution_def by auto
-qed
-
 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 finite_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
-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>'])
+  "random_variable MX X \<Longrightarrow> random_variable MY Y \<Longrightarrow> A \<in> sets MX \<Longrightarrow> B \<in> sets MY
+    \<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MX X) A"
+  by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets)
 
 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 (MX \<Otimes>\<^isub>M 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 (MX \<Otimes>\<^isub>M MY)" by default
-  have sa: "sigma_algebra M" by default
-  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
-    unfolding P.measurable_pair_iff[OF sa] using assms by (simp add: comp_def)
-qed
-
-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>'])
+  "random_variable MX X \<Longrightarrow> random_variable MY Y \<Longrightarrow> A \<in> sets MX \<Longrightarrow> B \<in> sets MY
+    \<Longrightarrow> emeasure (distr M (MX \<Otimes>\<^isub>M MY) (\<lambda>x. (X x, Y x))) (A \<times> B) \<le> emeasure (distr M MY Y) B"
+  by (auto simp: emeasure_distr measurable_pair_iff comp_def intro!: emeasure_mono measurable_sets)
 
 locale pair_prob_space = pair_sigma_finite M1 M2 + M1: prob_space M1 + M2: prob_space M2 for M1 M2
 
-sublocale pair_prob_space \<subseteq> P: prob_space P
+sublocale pair_prob_space \<subseteq> P: prob_space "M1 \<Otimes>\<^isub>M M2"
 proof
-  show "measure_space P" ..
-  show "measure P (space P) = 1"
-    by (simp add: pair_measure_times M1.measure_space_1 M2.measure_space_1 space_pair_measure)
+  show "emeasure (M1 \<Otimes>\<^isub>M M2) (space (M1 \<Otimes>\<^isub>M M2)) = 1"
+    by (simp add: emeasure_pair_measure_Times M1.emeasure_space_1 M2.emeasure_space_1 space_pair_measure)
 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>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 ((MX \<Otimes>\<^isub>M MY) \<lparr> measure := ereal \<circ> joint_distribution X Y\<rparr>)"
-  using random_variable_pairI[OF assms] by (rule distribution_prob_space)
-
-locale product_prob_space = product_sigma_finite M for M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme" +
+locale product_prob_space = product_sigma_finite M for M :: "'i \<Rightarrow> 'a measure" +
   fixes I :: "'i set"
   assumes prob_space: "\<And>i. prob_space (M i)"
 
@@ -578,648 +421,401 @@
 
 sublocale finite_product_prob_space \<subseteq> prob_space "\<Pi>\<^isub>M i\<in>I. M i"
 proof
-  show "measure_space P" ..
-  show "measure P (space P) = 1"
-    by (simp add: measure_times M.measure_space_1 setprod_1)
+  show "emeasure (\<Pi>\<^isub>M i\<in>I. M i) (space (\<Pi>\<^isub>M i\<in>I. M i)) = 1"
+    by (simp add: measure_times M.emeasure_space_1 setprod_1 space_PiM)
 qed
 
 lemma (in finite_product_prob_space) prob_times:
   assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets (M i)"
   shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
 proof -
-  have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
-    using X by (intro finite_measure_eq[symmetric] in_P) auto
-  also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))"
+  have "ereal (measure (\<Pi>\<^isub>M i\<in>I. M i) (\<Pi>\<^isub>E i\<in>I. X i)) = emeasure (\<Pi>\<^isub>M i\<in>I. M i) (\<Pi>\<^isub>E i\<in>I. X i)"
+    using X by (simp add: emeasure_eq_measure)
+  also have "\<dots> = (\<Prod>i\<in>I. emeasure (M i) (X i))"
     using measure_times X by simp
-  also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
-    using X by (simp add: M.finite_measure_eq setprod_ereal)
-  finally show ?thesis by simp
-qed
-
-lemma (in prob_space) random_variable_restrict:
-  assumes I: "finite I"
-  assumes X: "\<And>i. i \<in> I \<Longrightarrow> random_variable (N i) (X i)"
-  shows "random_variable (Pi\<^isub>M I N) (\<lambda>x. \<lambda>i\<in>I. X i x)"
-proof
-  { fix i assume "i \<in> I"
-    with X interpret N: sigma_algebra "N i" by simp
-    have "sets (N i) \<subseteq> Pow (space (N i))" by (rule N.space_closed) }
-  note N_closed = this
-  then show "sigma_algebra (Pi\<^isub>M I N)"
-    by (simp add: product_algebra_def)
-       (intro sigma_algebra_sigma product_algebra_generator_sets_into_space)
-  show "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable M (Pi\<^isub>M I N)"
-    using X by (intro measurable_restrict[OF I N_closed]) auto
-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
-  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\<lparr>measure := ereal \<circ> distribution X\<rparr>)"
-proof -
-  interpret X: prob_space "MX\<lparr>measure := ereal \<circ> distribution X\<rparr>"
-    using assms[THEN finite_random_variableD] by (rule distribution_prob_space)
-  interpret MX: finite_sigma_algebra MX
-    using assms by auto
-  show ?thesis by default (simp_all add: MX.finite_space)
-qed
-
-lemma (in prob_space) simple_function_imp_finite_random_variable[simp, intro]:
-  assumes "simple_function M X"
-  shows "finite_random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = x \<rparr> X"
-    (is "finite_random_variable ?X _")
-proof (intro conjI)
-  have [simp]: "finite (X ` space M)" using assms unfolding simple_function_def by simp
-  interpret X: sigma_algebra ?X by (rule sigma_algebra_Pow)
-  show "finite_sigma_algebra ?X"
-    by default auto
-  show "X \<in> measurable M ?X"
-  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 M X"
-  shows "random_variable \<lparr> space = X`space M, sets = Pow (X`space M), \<dots> = ext \<rparr> X"
-  using simple_function_imp_finite_random_variable[OF assms, of ext]
-  by (auto dest!: finite_random_variableD)
-
-lemma (in prob_space) sum_over_space_real_distribution:
-  "simple_function M X \<Longrightarrow> (\<Sum>x\<in>X`space M. distribution X {x}) = 1"
-  unfolding distribution_def prob_space[symmetric]
-  by (subst 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 (MX \<Otimes>\<^isub>M 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 (MX \<Otimes>\<^isub>M MY)" ..
-  have sa: "sigma_algebra M" by default
-  show "(\<lambda>x. (X x, Y x)) \<in> measurable M (MX \<Otimes>\<^isub>M MY)"
-    unfolding P.measurable_pair_iff[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_measurable:
-  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 finite_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_measurable[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
-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:
-  fixes MX :: "('c, 'd) measure_space_scheme" and MY :: "('e, 'f) measure_space_scheme"
-  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 intro: antisym)
-
-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 finite_measure_finite_Union[symmetric])
-  interpret MX: finite_sigma_algebra MX using X by auto
-  show "finite (space MX)" using MX.finite_space .
-  let ?d = "\<lambda>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_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
-
-locale pair_finite_prob_space = pair_prob_space M1 M2 + pair_finite_space M1 M2 + M1: finite_prob_space M1 + M2: finite_prob_space M2 for M1 M2
-
-sublocale pair_finite_prob_space \<subseteq> finite_prob_space P by default
-
-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_funcset_insert_eq[simp]:
-  assumes "a \<notin> A"
-  shows "extensional (insert a A) \<inter> (insert a A \<rightarrow> B) = (\<Union>f \<in> extensional 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 A \<inter> (A \<rightarrow> B))"
-  using assms by induct (auto simp: extensional_funcset_insert_eq)
-
-lemma finite_PiE[simp, intro]:
-  assumes fin: "finite A" "\<And>i. i \<in> A \<Longrightarrow> finite (B i)"
-  shows "finite (Pi\<^isub>E A B)"
-proof -
-  have *: "(Pi\<^isub>E A B) \<subseteq> extensional A \<inter> (A \<rightarrow> (\<Union>i\<in>A. B i))" by auto
-  show ?thesis
-    using fin by (intro finite_subset[OF *] finite_extensional_funcset) auto
-qed
-
-locale finite_product_finite_prob_space = finite_product_prob_space M I for M I +
-  assumes finite_space: "\<And>i. finite_prob_space (M i)"
-
-sublocale finite_product_finite_prob_space \<subseteq> M!: finite_prob_space "M i" using finite_space .
-
-lemma (in finite_product_finite_prob_space) singleton_eq_product:
-  assumes x: "x \<in> space P" shows "{x} = (\<Pi>\<^isub>E i\<in>I. {x i})"
-proof (safe intro!: ext[of _ x])
-  fix y i assume *: "y \<in> (\<Pi> i\<in>I. {x i})" "y \<in> extensional I"
-  with x show "y i = x i"
-    by (cases "i \<in> I") (auto simp: extensional_def)
-qed (insert x, auto)
-
-sublocale finite_product_finite_prob_space \<subseteq> finite_prob_space "Pi\<^isub>M I M"
-proof
-  show "finite (space P)"
-    using finite_index M.finite_space by auto
-
-  { fix x assume "x \<in> space P"
-    with this[THEN singleton_eq_product]
-    have "{x} \<in> sets P"
-      by (auto intro!: in_P) }
-  note x_in_P = this
-
-  have "Pow (space P) \<subseteq> sets P"
-  proof
-    fix X assume "X \<in> Pow (space P)"
-    moreover then have "finite X"
-      using `finite (space P)` by (blast intro: finite_subset)
-    ultimately have "(\<Union>x\<in>X. {x}) \<in> sets P"
-      by (intro finite_UN x_in_P) auto
-    then show "X \<in> sets P" by simp
-  qed
-  with space_closed show [simp]: "sets P = Pow (space P)" ..
-qed
-
-lemma (in finite_product_finite_prob_space) measure_finite_times:
-  "(\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)) \<Longrightarrow> \<mu> (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.\<mu> i (X i))"
-  by (rule measure_times) simp
-
-lemma (in finite_product_finite_prob_space) measure_singleton_times:
-  assumes x: "x \<in> space P" shows "\<mu> {x} = (\<Prod>i\<in>I. M.\<mu> i {x i})"
-  unfolding singleton_eq_product[OF x] using x
-  by (intro measure_finite_times) auto
-
-lemma (in finite_product_finite_prob_space) prob_finite_times:
-  assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<subseteq> space (M i)"
-  shows "prob (\<Pi>\<^isub>E i\<in>I. X i) = (\<Prod>i\<in>I. M.prob i (X i))"
-proof -
-  have "ereal (\<mu>' (\<Pi>\<^isub>E i\<in>I. X i)) = \<mu> (\<Pi>\<^isub>E i\<in>I. X i)"
-    using X by (intro finite_measure_eq[symmetric] in_P) auto
-  also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i (X i))"
-    using measure_finite_times X by simp
-  also have "\<dots> = ereal (\<Prod>i\<in>I. M.\<mu>' i (X i))"
-    using X by (simp add: M.finite_measure_eq setprod_ereal)
+  also have "\<dots> = ereal (\<Prod>i\<in>I. measure (M i) (X i))"
+    using X by (simp add: M.emeasure_eq_measure setprod_ereal)
   finally show ?thesis by simp
 qed
 
-lemma (in finite_product_finite_prob_space) prob_singleton_times:
-  assumes x: "x \<in> space P"
-  shows "prob {x} = (\<Prod>i\<in>I. M.prob i {x i})"
-  unfolding singleton_eq_product[OF x] using x
-  by (intro prob_finite_times) auto
-
-lemma (in finite_product_finite_prob_space) prob_finite_product:
-  "A \<subseteq> space P \<Longrightarrow> prob A = (\<Sum>x\<in>A. \<Prod>i\<in>I. M.prob i {x i})"
-  by (auto simp add: finite_measure_singleton prob_singleton_times
-           simp del: space_product_algebra
-           intro!: setsum_cong prob_singleton_times)
+section {* Distributions *}
 
-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 ((MX \<Otimes>\<^isub>M MY)\<lparr> measure := ereal \<circ> joint_distribution X Y\<rparr>)"
-  by (intro distribution_finite_prob_space finite_random_variable_pairI X Y)
-
-lemma finite_prob_space_eq:
-  "finite_prob_space M \<longleftrightarrow> finite_measure_space M \<and> measure M (space M) = 1"
-  unfolding finite_prob_space_def finite_measure_space_def prob_space_def prob_space_axioms_def
-  by auto
-
-lemma (in finite_prob_space) sum_over_space_eq_1: "(\<Sum>x\<in>space M. \<mu> {x}) = 1"
-  using measure_space_1 sum_over_space by simp
+definition "distributed M N X f \<longleftrightarrow> distr M N X = density N f \<and> 
+  f \<in> borel_measurable N \<and> (AE x in N. 0 \<le> f x) \<and> X \<in> measurable M N"
 
-lemma (in finite_prob_space) joint_distribution_restriction_fst:
-  "joint_distribution X Y A \<le> distribution X (fst ` A)"
-  unfolding distribution_def
-proof (safe intro!: finite_measure_mono)
-  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
-  show "x \<in> X -` fst ` A"
-    by (auto intro!: image_eqI[OF _ *])
-qed (simp_all add: sets_eq_Pow)
-
-lemma (in finite_prob_space) joint_distribution_restriction_snd:
-  "joint_distribution X Y A \<le> distribution Y (snd ` A)"
-  unfolding distribution_def
-proof (safe intro!: finite_measure_mono)
-  fix x assume "x \<in> space M" and *: "(X x, Y x) \<in> A"
-  show "x \<in> Y -` snd ` A"
-    by (auto intro!: image_eqI[OF _ *])
-qed (simp_all add: sets_eq_Pow)
-
-lemma (in finite_prob_space) distribution_order:
-  shows "0 \<le> distribution X x'"
-  and "(distribution X x' \<noteq> 0) \<longleftrightarrow> (0 < distribution X x')"
-  and "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_restriction_fst[of X Y "{(x, y)}"]
-    joint_distribution_restriction_snd[of X Y "{(x, y)}"]
-  by (auto intro: antisym)
-
-lemma (in finite_prob_space) distribution_mono:
-  assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
-  shows "distribution X x \<le> distribution Y y"
-  unfolding distribution_def
-  using assms by (auto simp: sets_eq_Pow intro!: finite_measure_mono)
+lemma
+  shows distributed_distr_eq_density: "distributed M N X f \<Longrightarrow> distr M N X = density N f"
+    and distributed_measurable: "distributed M N X f \<Longrightarrow> X \<in> measurable M N"
+    and distributed_borel_measurable: "distributed M N X f \<Longrightarrow> f \<in> borel_measurable N"
+    and distributed_AE: "distributed M N X f \<Longrightarrow> (AE x in N. 0 \<le> f x)"
+  by (simp_all add: distributed_def)
 
-lemma (in finite_prob_space) distribution_mono_gt_0:
-  assumes gt_0: "0 < distribution X x"
-  assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
-  shows "0 < distribution Y y"
-  by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
-
-lemma (in finite_prob_space) sum_over_space_distrib:
-  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
-  unfolding distribution_def prob_space[symmetric] using finite_space
-  by (subst finite_measure_finite_Union[symmetric])
-     (auto simp add: disjoint_family_on_def sets_eq_Pow
-           intro!: arg_cong[where f=\<mu>'])
-
-lemma (in finite_prob_space) finite_sum_over_space_eq_1:
-  "(\<Sum>x\<in>space M. prob {x}) = 1"
-  using prob_space finite_space
-  by (subst (asm) finite_measure_finite_singleton) auto
-
-lemma (in prob_space) distribution_remove_const:
-  shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
-  and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
-  and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
-  and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
-  and "distribution (\<lambda>x. ()) {()} = 1"
-  by (auto intro!: arg_cong[where f=\<mu>'] simp: distribution_def prob_space[symmetric])
+lemma
+  shows distributed_real_measurable: "distributed M N X (\<lambda>x. ereal (f x)) \<Longrightarrow> f \<in> borel_measurable N"
+    and distributed_real_AE: "distributed M N X (\<lambda>x. ereal (f x)) \<Longrightarrow> (AE x in N. 0 \<le> f x)"
+  by (simp_all add: distributed_def borel_measurable_ereal_iff)
 
-lemma (in finite_prob_space) setsum_distribution_gen:
-  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
-  and "inj_on f (X`space M)"
-  shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
-  unfolding distribution_def assms
-  using finite_space assms
-  by (subst finite_measure_finite_Union[symmetric])
-     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
-      intro!: arg_cong[where f=prob])
-
-lemma (in finite_prob_space) setsum_distribution_cut:
-  "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}"
-  "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}"
-  "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}"
-  "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}"
-  "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
-  by (auto intro!: inj_onI setsum_distribution_gen)
-
-lemma (in finite_prob_space) uniform_prob:
-  assumes "x \<in> space M"
-  assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
-  shows "prob {x} = 1 / card (space M)"
+lemma distributed_count_space:
+  assumes X: "distributed M (count_space A) X P" and a: "a \<in> A" and A: "finite A"
+  shows "P a = emeasure M (X -` {a} \<inter> space M)"
 proof -
-  have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
-    using assms(2)[OF _ `x \<in> space M`] by blast
-  have "1 = prob (space M)"
-    using prob_space by auto
-  also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
-    using finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
-      sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
-      finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
-    by (auto simp add:setsum_restrict_set)
-  also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
-    using prob_x by auto
-  also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
-  finally have one: "1 = real (card (space M)) * prob {x}"
-    using real_eq_of_nat by auto
-  hence two: "real (card (space M)) \<noteq> 0" by fastforce
-  from one have three: "prob {x} \<noteq> 0" by fastforce
-  thus ?thesis using one two three divide_cancel_right
-    by (auto simp:field_simps)
+  have "emeasure M (X -` {a} \<inter> space M) = emeasure (distr M (count_space A) X) {a}"
+    using X a A by (simp add: distributed_measurable emeasure_distr)
+  also have "\<dots> = emeasure (density (count_space A) P) {a}"
+    using X by (simp add: distributed_distr_eq_density)
+  also have "\<dots> = (\<integral>\<^isup>+x. P a * indicator {a} x \<partial>count_space A)"
+    using X a by (auto simp add: emeasure_density distributed_def indicator_def intro!: positive_integral_cong)
+  also have "\<dots> = P a"
+    using X a by (subst positive_integral_cmult_indicator) (auto simp: distributed_def one_ereal_def[symmetric] AE_count_space)
+  finally show ?thesis ..
 qed
 
-lemma (in prob_space) prob_space_subalgebra:
-  assumes "sigma_algebra N" "sets N \<subseteq> sets M" "space N = space M"
-    and "\<And>A. A \<in> sets N \<Longrightarrow> measure N A = \<mu> A"
-  shows "prob_space N"
-proof
-  interpret N: measure_space N
-    by (rule measure_space_subalgebra[OF assms])
-  show "measure_space N" ..
-  show "measure N (space N) = 1"
-    using assms(4)[OF N.top] by (simp add: assms measure_space_1)
+lemma distributed_cong_density:
+  "(AE x in N. f x = g x) \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable N \<Longrightarrow>
+    distributed M N X f \<longleftrightarrow> distributed M N X g"
+  by (auto simp: distributed_def intro!: density_cong)
+
+lemma subdensity:
+  assumes T: "T \<in> measurable P Q"
+  assumes f: "distributed M P X f"
+  assumes g: "distributed M Q Y g"
+  assumes Y: "Y = T \<circ> X"
+  shows "AE x in P. g (T x) = 0 \<longrightarrow> f x = 0"
+proof -
+  have "{x\<in>space Q. g x = 0} \<in> null_sets (distr M Q (T \<circ> X))"
+    using g Y by (auto simp: null_sets_density_iff distributed_def)
+  also have "distr M Q (T \<circ> X) = distr (distr M P X) Q T"
+    using T f[THEN distributed_measurable] by (rule distr_distr[symmetric])
+  finally have "T -` {x\<in>space Q. g x = 0} \<inter> space P \<in> null_sets (distr M P X)"
+    using T by (subst (asm) null_sets_distr_iff) auto
+  also have "T -` {x\<in>space Q. g x = 0} \<inter> space P = {x\<in>space P. g (T x) = 0}"
+    using T by (auto dest: measurable_space)
+  finally show ?thesis
+    using f g by (auto simp add: null_sets_density_iff distributed_def)
 qed
 
-lemma (in prob_space) prob_space_of_restricted_space:
-  assumes "\<mu> A \<noteq> 0" "A \<in> sets M"
-  shows "prob_space (restricted_space A \<lparr>measure := \<lambda>S. \<mu> S / \<mu> A\<rparr>)"
-    (is "prob_space ?P")
+lemma subdensity_real:
+  fixes g :: "'a \<Rightarrow> real" and f :: "'b \<Rightarrow> real"
+  assumes T: "T \<in> measurable P Q"
+  assumes f: "distributed M P X f"
+  assumes g: "distributed M Q Y g"
+  assumes Y: "Y = T \<circ> X"
+  shows "AE x in P. g (T x) = 0 \<longrightarrow> f x = 0"
+  using subdensity[OF T, of M X "\<lambda>x. ereal (f x)" Y "\<lambda>x. ereal (g x)"] assms by auto
+
+lemma distributed_integral:
+  "distributed M N X f \<Longrightarrow> g \<in> borel_measurable N \<Longrightarrow> (\<integral>x. f x * g x \<partial>N) = (\<integral>x. g (X x) \<partial>M)"
+  by (auto simp: distributed_real_measurable distributed_real_AE distributed_measurable
+                 distributed_distr_eq_density[symmetric] integral_density[symmetric] integral_distr)
+  
+lemma distributed_transform_integral:
+  assumes Px: "distributed M N X Px"
+  assumes "distributed M P Y Py"
+  assumes Y: "Y = T \<circ> X" and T: "T \<in> measurable N P" and f: "f \<in> borel_measurable P"
+  shows "(\<integral>x. Py x * f x \<partial>P) = (\<integral>x. Px x * f (T x) \<partial>N)"
 proof -
-  interpret A: measure_space "restricted_space A"
-    using `A \<in> sets M` by (rule restricted_measure_space)
-  interpret A': sigma_algebra ?P
-    by (rule A.sigma_algebra_cong) auto
-  show "prob_space ?P"
-  proof
-    show "measure_space ?P"
-    proof
-      show "positive ?P (measure ?P)"
-      proof (simp add: positive_def, safe)
-        fix B assume "B \<in> events"
-        with real_measure[of "A \<inter> B"] real_measure[OF `A \<in> events`] `A \<in> sets M`
-        show "0 \<le> \<mu> (A \<inter> B) / \<mu> A" by (auto simp: Int)
-      qed
-      show "countably_additive ?P (measure ?P)"
-      proof (simp add: countably_additive_def, safe)
-        fix B and F :: "nat \<Rightarrow> 'a set"
-        assume F: "range F \<subseteq> op \<inter> A ` events" "disjoint_family F"
-        { fix i
-          from F have "F i \<in> op \<inter> A ` events" by auto
-          with `A \<in> events` have "F i \<in> events" by auto }
-        moreover then have "range F \<subseteq> events" by auto
-        moreover have "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S"
-          by (simp add: mult_commute divide_ereal_def)
-        moreover have "0 \<le> inverse (\<mu> A)"
-          using real_measure[OF `A \<in> events`] by auto
-        ultimately show "(\<Sum>i. \<mu> (F i) / \<mu> A) = \<mu> (\<Union>i. F i) / \<mu> A"
-          using measure_countably_additive[of F] F
-          by (auto simp: suminf_cmult_ereal)
-      qed
-    qed
-    show "measure ?P (space ?P) = 1"
-      using real_measure[OF `A \<in> events`] `\<mu> A \<noteq> 0` by auto
-  qed
-qed
-
-lemma finite_prob_spaceI:
-  assumes "finite (space M)" "sets M = Pow(space M)"
-    and 1: "measure M (space M) = 1" and "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> measure M {x}"
-    and add: "\<And>A B. A \<subseteq> space M \<Longrightarrow> measure M A = (\<Sum>x\<in>A. measure M {x})"
-  shows "finite_prob_space M"
-proof -
-  interpret finite_measure_space M
-  proof
-    show "measure M (space M) \<noteq> \<infinity>" using 1 by simp
-  qed fact+
-  show ?thesis by default fact
-qed
-
-lemma (in finite_prob_space) distribution_eq_setsum:
-  "distribution X A = (\<Sum>x\<in>A \<inter> X ` space M. distribution X {x})"
-proof -
-  have *: "X -` A \<inter> space M = (\<Union>x\<in>A \<inter> X ` space M. X -` {x} \<inter> space M)"
-    by auto
-  then show "distribution X A = (\<Sum>x\<in>A \<inter> X ` space M. distribution X {x})"
-    using finite_space unfolding distribution_def *
-    by (intro finite_measure_finite_Union)
-       (auto simp: disjoint_family_on_def)
-qed
-
-lemma (in finite_prob_space) distribution_eq_setsum_finite:
-  assumes "finite A"
-  shows "distribution X A = (\<Sum>x\<in>A. distribution X {x})"
-proof -
-  note distribution_eq_setsum[of X A]
-  also have "(\<Sum>x\<in>A \<inter> X ` space M. distribution X {x}) = (\<Sum>x\<in>A. distribution X {x})"
-  proof (intro setsum_mono_zero_cong_left ballI)
-    fix i assume "i\<in>A - A \<inter> X ` space M"
-    then have "X -` {i} \<inter> space M = {}" by auto
-    then show "distribution X {i} = 0"
-      by (simp add: distribution_def)
-  next
-    show "finite A" by fact
-  qed simp_all
+  have "(\<integral>x. Py x * f x \<partial>P) = (\<integral>x. f (Y x) \<partial>M)"
+    by (rule distributed_integral) fact+
+  also have "\<dots> = (\<integral>x. f (T (X x)) \<partial>M)"
+    using Y by simp
+  also have "\<dots> = (\<integral>x. Px x * f (T x) \<partial>N)"
+    using measurable_comp[OF T f] Px by (intro distributed_integral[symmetric]) (auto simp: comp_def)
   finally show ?thesis .
 qed
 
-lemma (in finite_prob_space) finite_measure_space:
-  fixes X :: "'a \<Rightarrow> 'x"
-  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M), measure = ereal \<circ> distribution X\<rparr>"
-    (is "finite_measure_space ?S")
-proof (rule finite_measure_spaceI, simp_all)
-  show "finite (X ` space M)" using finite_space by simp
-next
-  fix A assume "A \<subseteq> X ` space M"
-  then show "distribution X A = (\<Sum>x\<in>A. distribution X {x})"
-    by (subst distribution_eq_setsum) (simp add: Int_absorb2)
+lemma distributed_marginal_eq_joint:
+  assumes T: "sigma_finite_measure T"
+  assumes S: "sigma_finite_measure S"
+  assumes Px: "distributed M S X Px"
+  assumes Py: "distributed M T Y Py"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  shows "AE y in T. Py y = (\<integral>\<^isup>+x. Pxy (x, y) \<partial>S)"
+proof (rule sigma_finite_measure.density_unique[OF T])
+  interpret ST: pair_sigma_finite S T using S T unfolding pair_sigma_finite_def by simp
+  show "Py \<in> borel_measurable T" "AE y in T. 0 \<le> Py y"
+    "(\<lambda>x. \<integral>\<^isup>+ xa. Pxy (xa, x) \<partial>S) \<in> borel_measurable T" "AE y in T. 0 \<le> \<integral>\<^isup>+ x. Pxy (x, y) \<partial>S"
+    using Pxy[THEN distributed_borel_measurable]
+    by (auto intro!: Py[THEN distributed_borel_measurable] Py[THEN distributed_AE]
+                     ST.positive_integral_snd_measurable' positive_integral_positive)
+
+  show "density T Py = density T (\<lambda>x. \<integral>\<^isup>+ xa. Pxy (xa, x) \<partial>S)"
+  proof (rule measure_eqI)
+    fix A assume A: "A \<in> sets (density T Py)"
+    have *: "\<And>x y. x \<in> space S \<Longrightarrow> indicator (space S \<times> A) (x, y) = indicator A y"
+      by (auto simp: indicator_def)
+    have "emeasure (density T Py) A = emeasure (distr M T Y) A"
+      unfolding Py[THEN distributed_distr_eq_density] ..
+    also have "\<dots> = emeasure (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) (space S \<times> A)"
+      using A Px Py Pxy
+      by (subst (1 2) emeasure_distr)
+         (auto dest: measurable_space distributed_measurable intro!: arg_cong[where f="emeasure M"])
+    also have "\<dots> = emeasure (density (S \<Otimes>\<^isub>M T) Pxy) (space S \<times> A)"
+      unfolding Pxy[THEN distributed_distr_eq_density] ..
+    also have "\<dots> = (\<integral>\<^isup>+ x. Pxy x * indicator (space S \<times> A) x \<partial>(S \<Otimes>\<^isub>M T))"
+      using A Pxy by (simp add: emeasure_density distributed_borel_measurable)
+    also have "\<dots> = (\<integral>\<^isup>+y. \<integral>\<^isup>+x. Pxy (x, y) * indicator (space S \<times> A) (x, y) \<partial>S \<partial>T)"
+      using A Pxy
+      by (subst ST.positive_integral_snd_measurable) (simp_all add: emeasure_density distributed_borel_measurable)
+    also have "\<dots> = (\<integral>\<^isup>+y. (\<integral>\<^isup>+x. Pxy (x, y) \<partial>S) * indicator A y \<partial>T)"
+      using measurable_comp[OF measurable_Pair1[OF measurable_identity] distributed_borel_measurable[OF Pxy]]
+      using distributed_borel_measurable[OF Pxy] distributed_AE[OF Pxy, THEN ST.AE_pair]
+      by (subst (asm) ST.AE_commute) (auto intro!: positive_integral_cong_AE positive_integral_multc cong: positive_integral_cong simp: * comp_def)
+    also have "\<dots> = emeasure (density T (\<lambda>x. \<integral>\<^isup>+ xa. Pxy (xa, x) \<partial>S)) A"
+      using A by (intro emeasure_density[symmetric])  (auto intro!: ST.positive_integral_snd_measurable' Pxy[THEN distributed_borel_measurable])
+    finally show "emeasure (density T Py) A = emeasure (density T (\<lambda>x. \<integral>\<^isup>+ xa. Pxy (xa, x) \<partial>S)) A" .
+  qed simp
 qed
 
-lemma (in finite_prob_space) finite_prob_space_of_images:
-  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M), measure = ereal \<circ> distribution X \<rparr>"
-  by (simp add: finite_prob_space_eq finite_measure_space measure_space_1 one_ereal_def)
+lemma (in prob_space) distr_marginal1:
+  fixes Pxy :: "('b \<times> 'c) \<Rightarrow> real"
+  assumes "sigma_finite_measure S" "sigma_finite_measure T"
+  assumes Pxy: "distributed M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x)) Pxy"
+  defines "Px \<equiv> \<lambda>x. real (\<integral>\<^isup>+z. Pxy (x, z) \<partial>T)"
+  shows "distributed M S X Px"
+  unfolding distributed_def
+proof safe
+  interpret S: sigma_finite_measure S by fact
+  interpret T: sigma_finite_measure T by fact
+  interpret ST: pair_sigma_finite S T by default
+
+  have XY: "(\<lambda>x. (X x, Y x)) \<in> measurable M (S \<Otimes>\<^isub>M T)"
+    using Pxy by (rule distributed_measurable)
+  then show X: "X \<in> measurable M S"
+    unfolding measurable_pair_iff by (simp add: comp_def)
+  from XY have Y: "Y \<in> measurable M T"
+    unfolding measurable_pair_iff by (simp add: comp_def)
+
+  from Pxy show borel: "(\<lambda>x. ereal (Px x)) \<in> borel_measurable S"
+    by (auto intro!: ST.positive_integral_fst_measurable borel_measurable_real_of_ereal dest!: distributed_real_measurable simp: Px_def)
 
-lemma (in finite_prob_space) finite_product_measure_space:
-  fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
-  assumes "finite s1" "finite s2"
-  shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2), measure = ereal \<circ> joint_distribution X Y\<rparr>"
-    (is "finite_measure_space ?M")
-proof (rule finite_measure_spaceI, simp_all)
-  show "finite (s1 \<times> s2)"
-    using assms by auto
-next
-  fix A assume "A \<subseteq> (s1 \<times> s2)"
-  with assms show "joint_distribution X Y A = (\<Sum>x\<in>A. joint_distribution X Y {x})"
-    by (intro distribution_eq_setsum_finite) (auto dest: finite_subset)
-qed
-
-lemma (in finite_prob_space) finite_product_measure_space_of_images:
-  shows "finite_measure_space \<lparr> space = X ` space M \<times> Y ` space M,
-                                sets = Pow (X ` space M \<times> Y ` space M),
-                                measure = ereal \<circ> joint_distribution X Y \<rparr>"
-  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),
-                       measure = ereal \<circ> joint_distribution X Y \<rparr>"
-  (is "finite_prob_space ?S")
-proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images one_ereal_def)
-  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)
+  interpret Pxy: prob_space "distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))"
+    using XY by (rule prob_space_distr)
+  have "(\<integral>\<^isup>+ x. max 0 (ereal (- Pxy x)) \<partial>(S \<Otimes>\<^isub>M T)) = (\<integral>\<^isup>+ x. 0 \<partial>(S \<Otimes>\<^isub>M T))"
+    using Pxy
+    by (intro positive_integral_cong_AE) (auto simp: max_def dest: distributed_real_measurable distributed_real_AE)
+  then have Pxy_integrable: "integrable (S \<Otimes>\<^isub>M T) Pxy"
+    using Pxy Pxy.emeasure_space_1
+    by (simp add: integrable_def emeasure_density positive_integral_max_0 distributed_def borel_measurable_ereal_iff cong: positive_integral_cong)
+    
+  show "distr M S X = density S Px"
+  proof (rule measure_eqI)
+    fix A assume A: "A \<in> sets (distr M S X)"
+    with X Y XY have "emeasure (distr M S X) A = emeasure (distr M (S \<Otimes>\<^isub>M T) (\<lambda>x. (X x, Y x))) (A \<times> space T)"
+      by (auto simp add: emeasure_distr
+               intro!: arg_cong[where f="emeasure M"] dest: measurable_space)
+    also have "\<dots> = emeasure (density (S \<Otimes>\<^isub>M T) Pxy) (A \<times> space T)"
+      using Pxy by (simp add: distributed_def)
+    also have "\<dots> = \<integral>\<^isup>+ x. \<integral>\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \<times> space T) (x, y) \<partial>T \<partial>S"
+      using A borel Pxy
+      by (simp add: emeasure_density ST.positive_integral_fst_measurable(2)[symmetric] distributed_def)
+    also have "\<dots> = \<integral>\<^isup>+ x. ereal (Px x) * indicator A x \<partial>S"
+      apply (rule positive_integral_cong_AE)
+      using Pxy[THEN distributed_real_AE, THEN ST.AE_pair] ST.integrable_fst_measurable(1)[OF Pxy_integrable] AE_space
+    proof eventually_elim
+      fix x assume "x \<in> space S" "AE y in T. 0 \<le> Pxy (x, y)" and i: "integrable T (\<lambda>y. Pxy (x, y))"
+      moreover have eq: "\<And>y. y \<in> space T \<Longrightarrow> indicator (A \<times> space T) (x, y) = indicator A x"
+        by (auto simp: indicator_def)
+      ultimately have "(\<integral>\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \<times> space T) (x, y) \<partial>T) =
+          (\<integral>\<^isup>+ y. ereal (Pxy (x, y)) \<partial>T) * indicator A x"
+        using Pxy[THEN distributed_real_measurable] by (simp add: eq positive_integral_multc measurable_Pair2 cong: positive_integral_cong)
+      also have "(\<integral>\<^isup>+ y. ereal (Pxy (x, y)) \<partial>T) = Px x"
+        using i by (simp add: Px_def ereal_real integrable_def positive_integral_positive)
+      finally show "(\<integral>\<^isup>+ y. ereal (Pxy (x, y)) * indicator (A \<times> space T) (x, y) \<partial>T) = ereal (Px x) * indicator A x" .
+    qed
+    finally show "emeasure (distr M S X) A = emeasure (density S Px) A"
+      using A borel Pxy by (simp add: emeasure_density)
+  qed simp
+  
+  show "AE x in S. 0 \<le> ereal (Px x)"
+    by (simp add: Px_def positive_integral_positive real_of_ereal_pos)
 qed
 
-subsection "Borel Measure on {0 ..< 1}"
-
-definition pborel :: "real measure_space" where
-  "pborel = lborel.restricted_space {0 ..< 1}"
-
-lemma space_pborel[simp]:
-  "space pborel = {0 ..< 1}"
-  unfolding pborel_def by auto
-
-lemma sets_pborel:
-  "A \<in> sets pborel \<longleftrightarrow> A \<in> sets borel \<and> A \<subseteq> { 0 ..< 1}"
-  unfolding pborel_def by auto
+definition
+  "simple_distributed M X f \<longleftrightarrow> distributed M (count_space (X`space M)) X (\<lambda>x. ereal (f x)) \<and>
+    finite (X`space M)"
 
-lemma in_pborel[intro, simp]:
-  "A \<subseteq> {0 ..< 1} \<Longrightarrow> A \<in> sets borel \<Longrightarrow> A \<in> sets pborel"
-  unfolding pborel_def by auto
-
-interpretation pborel: measure_space pborel
-  using lborel.restricted_measure_space[of "{0 ..< 1}"]
-  by (simp add: pborel_def)
+lemma simple_distributed:
+  "simple_distributed M X Px \<Longrightarrow> distributed M (count_space (X`space M)) X Px"
+  unfolding simple_distributed_def by auto
 
-interpretation pborel: prob_space pborel
-proof
-  show "measure pborel (space pborel) = 1"
-    by (simp add: one_ereal_def pborel_def)
-qed default
-
-lemma pborel_prob: "pborel.prob A = (if A \<in> sets borel \<and> A \<subseteq> {0 ..< 1} then real (lborel.\<mu> A) else 0)"
-  unfolding pborel.\<mu>'_def by (auto simp: pborel_def)
+lemma simple_distributed_finite[dest]: "simple_distributed M X P \<Longrightarrow> finite (X`space M)"
+  by (simp add: simple_distributed_def)
 
-lemma pborel_singelton[simp]: "pborel.prob {a} = 0"
-  by (auto simp: pborel_prob)
-
-lemma
-  shows pborel_atLeastAtMost[simp]: "pborel.\<mu>' {a .. b} = (if 0 \<le> a \<and> a \<le> b \<and> b < 1 then b - a else 0)"
-    and pborel_atLeastLessThan[simp]: "pborel.\<mu>' {a ..< b} = (if 0 \<le> a \<and> a \<le> b \<and> b \<le> 1 then b - a else 0)"
-    and pborel_greaterThanAtMost[simp]: "pborel.\<mu>' {a <.. b} = (if 0 \<le> a \<and> a \<le> b \<and> b < 1 then b - a else 0)"
-    and pborel_greaterThanLessThan[simp]: "pborel.\<mu>' {a <..< b} = (if 0 \<le> a \<and> a \<le> b \<and> b \<le> 1 then b - a else 0)"
-  unfolding pborel_prob
-  by (auto simp: atLeastAtMost_subseteq_atLeastLessThan_iff
-    greaterThanAtMost_subseteq_atLeastLessThan_iff greaterThanLessThan_subseteq_atLeastLessThan_iff)
-
-lemma pborel_lebesgue_measure:
-  "A \<in> sets pborel \<Longrightarrow> pborel.prob A = real (measure lebesgue A)"
-  by (simp add: sets_pborel pborel_prob)
+lemma (in prob_space) distributed_simple_function_superset:
+  assumes X: "simple_function M X" "\<And>x. x \<in> X ` space M \<Longrightarrow> P x = measure M (X -` {x} \<inter> space M)"
+  assumes A: "X`space M \<subseteq> A" "finite A"
+  defines "S \<equiv> count_space A" and "P' \<equiv> (\<lambda>x. if x \<in> X`space M then P x else 0)"
+  shows "distributed M S X P'"
+  unfolding distributed_def
+proof safe
+  show "(\<lambda>x. ereal (P' x)) \<in> borel_measurable S" unfolding S_def by simp
+  show "AE x in S. 0 \<le> ereal (P' x)"
+    using X by (auto simp: S_def P'_def simple_distributed_def intro!: measure_nonneg)
+  show "distr M S X = density S P'"
+  proof (rule measure_eqI_finite)
+    show "sets (distr M S X) = Pow A" "sets (density S P') = Pow A"
+      using A unfolding S_def by auto
+    show "finite A" by fact
+    fix a assume a: "a \<in> A"
+    then have "a \<notin> X`space M \<Longrightarrow> X -` {a} \<inter> space M = {}" by auto
+    with A a X have "emeasure (distr M S X) {a} = P' a"
+      by (subst emeasure_distr)
+         (auto simp add: S_def P'_def simple_functionD emeasure_eq_measure
+               intro!: arg_cong[where f=prob])
+    also have "\<dots> = (\<integral>\<^isup>+x. ereal (P' a) * indicator {a} x \<partial>S)"
+      using A X a
+      by (subst positive_integral_cmult_indicator)
+         (auto simp: S_def P'_def simple_distributed_def simple_functionD measure_nonneg)
+    also have "\<dots> = (\<integral>\<^isup>+x. ereal (P' x) * indicator {a} x \<partial>S)"
+      by (auto simp: indicator_def intro!: positive_integral_cong)
+    also have "\<dots> = emeasure (density S P') {a}"
+      using a A by (intro emeasure_density[symmetric]) (auto simp: S_def)
+    finally show "emeasure (distr M S X) {a} = emeasure (density S P') {a}" .
+  qed
+  show "random_variable S X"
+    using X(1) A by (auto simp: measurable_def simple_functionD S_def)
+qed
 
-lemma pborel_alt:
-  "pborel = sigma \<lparr>
-    space = {0..<1},
-    sets = range (\<lambda>(x,y). {x..<y} \<inter> {0..<1}),
-    measure = measure lborel \<rparr>" (is "_ = ?R")
+lemma (in prob_space) simple_distributedI:
+  assumes X: "simple_function M X" "\<And>x. x \<in> X ` space M \<Longrightarrow> P x = measure M (X -` {x} \<inter> space M)"
+  shows "simple_distributed M X P"
+  unfolding simple_distributed_def
+proof
+  have "distributed M (count_space (X ` space M)) X (\<lambda>x. ereal (if x \<in> X`space M then P x else 0))"
+    (is "?A")
+    using simple_functionD[OF X(1)] by (intro distributed_simple_function_superset[OF X]) auto
+  also have "?A \<longleftrightarrow> distributed M (count_space (X ` space M)) X (\<lambda>x. ereal (P x))"
+    by (rule distributed_cong_density) auto
+  finally show "\<dots>" .
+qed (rule simple_functionD[OF X(1)])
+
+lemma simple_distributed_joint_finite:
+  assumes X: "simple_distributed M (\<lambda>x. (X x, Y x)) Px"
+  shows "finite (X ` space M)" "finite (Y ` space M)"
 proof -
-  have *: "{0..<1::real} \<in> sets borel" by auto
-  have **: "op \<inter> {0..<1::real} ` range (\<lambda>(x, y). {x..<y}) = range (\<lambda>(x,y). {x..<y} \<inter> {0..<1})"
-    unfolding image_image by (intro arg_cong[where f=range]) auto
-  have "pborel = algebra.restricted_space (sigma \<lparr>space=UNIV, sets=range (\<lambda>(a, b). {a ..< b :: real}),
-    measure = measure pborel\<rparr>) {0 ..< 1}"
-    by (simp add: sigma_def lebesgue_def pborel_def borel_eq_atLeastLessThan lborel_def)
-  also have "\<dots> = ?R"
-    by (subst restricted_sigma)
-       (simp_all add: sets_sigma sigma_sets.Basic ** pborel_def image_eqI[of _ _ "(0,1)"])
-  finally show ?thesis .
+  have "finite ((\<lambda>x. (X x, Y x)) ` space M)"
+    using X by (auto simp: simple_distributed_def simple_functionD)
+  then have "finite (fst ` (\<lambda>x. (X x, Y x)) ` space M)" "finite (snd ` (\<lambda>x. (X x, Y x)) ` space M)"
+    by auto
+  then show fin: "finite (X ` space M)" "finite (Y ` space M)"
+    by (auto simp: image_image)
+qed
+
+lemma simple_distributed_joint2_finite:
+  assumes X: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Px"
+  shows "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)"
+proof -
+  have "finite ((\<lambda>x. (X x, Y x, Z x)) ` space M)"
+    using X by (auto simp: simple_distributed_def simple_functionD)
+  then have "finite (fst ` (\<lambda>x. (X x, Y x, Z x)) ` space M)"
+    "finite ((fst \<circ> snd) ` (\<lambda>x. (X x, Y x, Z x)) ` space M)"
+    "finite ((snd \<circ> snd) ` (\<lambda>x. (X x, Y x, Z x)) ` space M)"
+    by auto
+  then show fin: "finite (X ` space M)" "finite (Y ` space M)" "finite (Z ` space M)"
+    by (auto simp: image_image)
 qed
 
-subsection "Bernoulli space"
+lemma simple_distributed_simple_function:
+  "simple_distributed M X Px \<Longrightarrow> simple_function M X"
+  unfolding simple_distributed_def distributed_def
+  by (auto simp: simple_function_def)
+
+lemma simple_distributed_measure:
+  "simple_distributed M X P \<Longrightarrow> a \<in> X`space M \<Longrightarrow> P a = measure M (X -` {a} \<inter> space M)"
+  using distributed_count_space[of M "X`space M" X P a, symmetric]
+  by (auto simp: simple_distributed_def measure_def)
+
+lemma simple_distributed_nonneg: "simple_distributed M X f \<Longrightarrow> x \<in> space M \<Longrightarrow> 0 \<le> f (X x)"
+  by (auto simp: simple_distributed_measure measure_nonneg)
 
-definition "bernoulli_space p = \<lparr> space = UNIV, sets = UNIV,
-  measure = ereal \<circ> setsum (\<lambda>b. if b then min 1 (max 0 p) else 1 - min 1 (max 0 p)) \<rparr>"
+lemma (in prob_space) simple_distributed_joint:
+  assumes X: "simple_distributed M (\<lambda>x. (X x, Y x)) Px"
+  defines "S \<equiv> count_space (X`space M) \<Otimes>\<^isub>M count_space (Y`space M)"
+  defines "P \<equiv> (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x))`space M then Px x else 0)"
+  shows "distributed M S (\<lambda>x. (X x, Y x)) P"
+proof -
+  from simple_distributed_joint_finite[OF X, simp]
+  have S_eq: "S = count_space (X`space M \<times> Y`space M)"
+    by (simp add: S_def pair_measure_count_space)
+  show ?thesis
+    unfolding S_eq P_def
+  proof (rule distributed_simple_function_superset)
+    show "simple_function M (\<lambda>x. (X x, Y x))"
+      using X by (rule simple_distributed_simple_function)
+    fix x assume "x \<in> (\<lambda>x. (X x, Y x)) ` space M"
+    from simple_distributed_measure[OF X this]
+    show "Px x = prob ((\<lambda>x. (X x, Y x)) -` {x} \<inter> space M)" .
+  qed auto
+qed
 
-interpretation bernoulli: finite_prob_space "bernoulli_space p" for p
-  by (rule finite_prob_spaceI)
-     (auto simp: bernoulli_space_def UNIV_bool one_ereal_def setsum_Un_disjoint intro!: setsum_nonneg)
+lemma (in prob_space) simple_distributed_joint2:
+  assumes X: "simple_distributed M (\<lambda>x. (X x, Y x, Z x)) Px"
+  defines "S \<equiv> count_space (X`space M) \<Otimes>\<^isub>M count_space (Y`space M) \<Otimes>\<^isub>M count_space (Z`space M)"
+  defines "P \<equiv> (\<lambda>x. if x \<in> (\<lambda>x. (X x, Y x, Z x))`space M then Px x else 0)"
+  shows "distributed M S (\<lambda>x. (X x, Y x, Z x)) P"
+proof -
+  from simple_distributed_joint2_finite[OF X, simp]
+  have S_eq: "S = count_space (X`space M \<times> Y`space M \<times> Z`space M)"
+    by (simp add: S_def pair_measure_count_space)
+  show ?thesis
+    unfolding S_eq P_def
+  proof (rule distributed_simple_function_superset)
+    show "simple_function M (\<lambda>x. (X x, Y x, Z x))"
+      using X by (rule simple_distributed_simple_function)
+    fix x assume "x \<in> (\<lambda>x. (X x, Y x, Z x)) ` space M"
+    from simple_distributed_measure[OF X this]
+    show "Px x = prob ((\<lambda>x. (X x, Y x, Z x)) -` {x} \<inter> space M)" .
+  qed auto
+qed
+
+lemma (in prob_space) simple_distributed_setsum_space:
+  assumes X: "simple_distributed M X f"
+  shows "setsum f (X`space M) = 1"
+proof -
+  from X have "setsum f (X`space M) = prob (\<Union>i\<in>X`space M. X -` {i} \<inter> space M)"
+    by (subst finite_measure_finite_Union)
+       (auto simp add: disjoint_family_on_def simple_distributed_measure simple_distributed_simple_function simple_functionD
+             intro!: setsum_cong arg_cong[where f="prob"])
+  also have "\<dots> = prob (space M)"
+    by (auto intro!: arg_cong[where f=prob])
+  finally show ?thesis
+    using emeasure_space_1 by (simp add: emeasure_eq_measure one_ereal_def)
+qed
 
-lemma bernoulli_measure:
-  "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p B = (\<Sum>b\<in>B. if b then p else 1 - p)"
-  unfolding bernoulli.\<mu>'_def unfolding bernoulli_space_def by (auto intro!: setsum_cong)
+lemma (in prob_space) distributed_marginal_eq_joint_simple:
+  assumes Px: "simple_function M X"
+  assumes Py: "simple_distributed M Y Py"
+  assumes Pxy: "simple_distributed M (\<lambda>x. (X x, Y x)) Pxy"
+  assumes y: "y \<in> Y`space M"
+  shows "Py y = (\<Sum>x\<in>X`space M. if (x, y) \<in> (\<lambda>x. (X x, Y x)) ` space M then Pxy (x, y) else 0)"
+proof -
+  note Px = simple_distributedI[OF Px refl]
+  have *: "\<And>f A. setsum (\<lambda>x. max 0 (ereal (f x))) A = ereal (setsum (\<lambda>x. max 0 (f x)) A)"
+    by (simp add: setsum_ereal[symmetric] zero_ereal_def)
+  from distributed_marginal_eq_joint[OF sigma_finite_measure_count_space_finite sigma_finite_measure_count_space_finite
+    simple_distributed[OF Px] simple_distributed[OF Py] simple_distributed_joint[OF Pxy],
+    OF Py[THEN simple_distributed_finite] Px[THEN simple_distributed_finite]]
+    y Px[THEN simple_distributed_finite] Py[THEN simple_distributed_finite]
+    Pxy[THEN simple_distributed, THEN distributed_real_AE]
+  show ?thesis
+    unfolding AE_count_space
+    apply (elim ballE[where x=y])
+    apply (auto simp add: positive_integral_count_space_finite * intro!: setsum_cong split: split_max)
+    done
+qed
 
-lemma bernoulli_measure_True: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {True} = p"
-  and bernoulli_measure_False: "0 \<le> p \<Longrightarrow> p \<le> 1 \<Longrightarrow> bernoulli.prob p {False} = 1 - p"
-  unfolding bernoulli_measure by simp_all
+
+lemma prob_space_uniform_measure:
+  assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>"
+  shows "prob_space (uniform_measure M A)"
+proof
+  show "emeasure (uniform_measure M A) (space (uniform_measure M A)) = 1"
+    using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)], of "space M"]
+    using sets_into_space[OF emeasure_neq_0_sets[OF A(1)]] A
+    by (simp add: Int_absorb2 emeasure_nonneg)
+qed
+
+lemma prob_space_uniform_count_measure: "finite A \<Longrightarrow> A \<noteq> {} \<Longrightarrow> prob_space (uniform_count_measure A)"
+  by default (auto simp: emeasure_uniform_count_measure space_uniform_count_measure one_ereal_def)
 
 end
--- a/src/HOL/Probability/Radon_Nikodym.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -8,45 +8,79 @@
 imports Lebesgue_Integration
 begin
 
+definition "diff_measure M N =
+  measure_of (space M) (sets M) (\<lambda>A. emeasure M A - emeasure N A)"
+
+lemma 
+  shows space_diff_measure[simp]: "space (diff_measure M N) = space M"
+    and sets_diff_measure[simp]: "sets (diff_measure M N) = sets M"
+  by (auto simp: diff_measure_def)
+
+lemma emeasure_diff_measure:
+  assumes fin: "finite_measure M" "finite_measure N" and sets_eq: "sets M = sets N"
+  assumes pos: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure N A \<le> emeasure M A" and A: "A \<in> sets M"
+  shows "emeasure (diff_measure M N) A = emeasure M A - emeasure N A" (is "_ = ?\<mu> A")
+  unfolding diff_measure_def
+proof (rule emeasure_measure_of_sigma)
+  show "sigma_algebra (space M) (sets M)" ..
+  show "positive (sets M) ?\<mu>"
+    using pos by (simp add: positive_def ereal_diff_positive)
+  show "countably_additive (sets M) ?\<mu>"
+  proof (rule countably_additiveI)
+    fix A :: "nat \<Rightarrow> _"  assume A: "range A \<subseteq> sets M" and "disjoint_family A"
+    then have suminf:
+      "(\<Sum>i. emeasure M (A i)) = emeasure M (\<Union>i. A i)"
+      "(\<Sum>i. emeasure N (A i)) = emeasure N (\<Union>i. A i)"
+      by (simp_all add: suminf_emeasure sets_eq)
+    with A have "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
+      (\<Sum>i. emeasure M (A i)) - (\<Sum>i. emeasure N (A i))"
+      using fin
+      by (intro suminf_ereal_minus pos emeasure_nonneg)
+         (auto simp: sets_eq finite_measure.emeasure_eq_measure suminf_emeasure)
+    then show "(\<Sum>i. emeasure M (A i) - emeasure N (A i)) =
+      emeasure M (\<Union>i. A i) - emeasure N (\<Union>i. A i) "
+      by (simp add: suminf)
+  qed
+qed fact
+
 lemma (in sigma_finite_measure) Ex_finite_integrable_function:
   shows "\<exists>h\<in>borel_measurable M. integral\<^isup>P M h \<noteq> \<infinity> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>) \<and> (\<forall>x. 0 \<le> h x)"
 proof -
   obtain A :: "nat \<Rightarrow> 'a set" where
     range: "range A \<subseteq> sets M" and
     space: "(\<Union>i. A i) = space M" and
-    measure: "\<And>i. \<mu> (A i) \<noteq> \<infinity>" and
+    measure: "\<And>i. emeasure M (A i) \<noteq> \<infinity>" and
     disjoint: "disjoint_family A"
-    using disjoint_sigma_finite by auto
-  let ?B = "\<lambda>i. 2^Suc i * \<mu> (A i)"
+    using sigma_finite_disjoint by auto
+  let ?B = "\<lambda>i. 2^Suc i * emeasure M (A i)"
   have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
   proof
-    fix i have Ai: "A i \<in> sets M" using range by auto
-    from measure positive_measure[OF this]
-    show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
-      by (auto intro!: ereal_dense simp: ereal_0_gt_inverse)
+    fix i show "\<exists>x. 0 < x \<and> x < inverse (?B i)"
+      using measure[of i] emeasure_nonneg[of M "A i"]
+      by (auto intro!: ereal_dense simp: ereal_0_gt_inverse ereal_zero_le_0_iff)
   qed
   from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
-    "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
+    "\<And>i. n i < inverse (2^Suc i * emeasure M (A i))" by auto
   { fix i have "0 \<le> n i" using n(1)[of i] by auto } note pos = this
   let ?h = "\<lambda>x. \<Sum>i. n i * indicator (A i) x"
   show ?thesis
   proof (safe intro!: bexI[of _ ?h] del: notI)
     have "\<And>i. A i \<in> sets M"
       using range by fastforce+
-    then have "integral\<^isup>P M ?h = (\<Sum>i. n i * \<mu> (A i))" using pos
+    then have "integral\<^isup>P M ?h = (\<Sum>i. n i * emeasure M (A i))" using pos
       by (simp add: positive_integral_suminf positive_integral_cmult_indicator)
     also have "\<dots> \<le> (\<Sum>i. (1 / 2)^Suc i)"
     proof (rule suminf_le_pos)
       fix N
-      have "n N * \<mu> (A N) \<le> inverse (2^Suc N * \<mu> (A N)) * \<mu> (A N)"
-        using positive_measure[OF `A N \<in> sets M`] n[of N]
+      have "n N * emeasure M (A N) \<le> inverse (2^Suc N * emeasure M (A N)) * emeasure M (A N)"
+        using n[of N]
         by (intro ereal_mult_right_mono) auto
       also have "\<dots> \<le> (1 / 2) ^ Suc N"
         using measure[of N] n[of N]
-        by (cases rule: ereal2_cases[of "n N" "\<mu> (A N)"])
+        by (cases rule: ereal2_cases[of "n N" "emeasure M (A N)"])
            (simp_all add: inverse_eq_divide power_divide one_ereal_def ereal_power_divide)
-      finally show "n N * \<mu> (A N) \<le> (1 / 2) ^ Suc N" .
-      show "0 \<le> n N * \<mu> (A N)" using n[of N] `A N \<in> sets M` by simp
+      finally show "n N * emeasure M (A N) \<le> (1 / 2) ^ Suc N" .
+      show "0 \<le> n N * emeasure M (A N)" using n[of N] `A N \<in> sets M` by (simp add: emeasure_nonneg)
     qed
     finally show "integral\<^isup>P M ?h \<noteq> \<infinity>" unfolding suminf_half_series_ereal by auto
   next
@@ -71,68 +105,45 @@
 
 subsection "Absolutely continuous"
 
-definition (in measure_space)
-  "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: ereal))"
+definition absolutely_continuous :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> bool" where
+  "absolutely_continuous M N \<longleftrightarrow> null_sets M \<subseteq> null_sets N"
+
+lemma absolutely_continuousI_count_space: "absolutely_continuous (count_space A) M"
+  unfolding absolutely_continuous_def by (auto simp: null_sets_count_space)
 
-lemma (in measure_space) absolutely_continuous_AE:
-  assumes "measure_space M'" and [simp]: "sets M' = sets M" "space M' = space M"
-    and "absolutely_continuous (measure M')" "AE x. P x"
+lemma absolutely_continuousI_density:
+  "f \<in> borel_measurable M \<Longrightarrow> absolutely_continuous M (density M f)"
+  by (force simp add: absolutely_continuous_def null_sets_density_iff dest: AE_not_in)
+
+lemma absolutely_continuousI_point_measure_finite:
+  "(\<And>x. \<lbrakk> x \<in> A ; f x \<le> 0 \<rbrakk> \<Longrightarrow> g x \<le> 0) \<Longrightarrow> absolutely_continuous (point_measure A f) (point_measure A g)"
+  unfolding absolutely_continuous_def by (force simp: null_sets_point_measure_iff)
+
+lemma absolutely_continuous_AE:
+  assumes sets_eq: "sets M' = sets M"
+    and "absolutely_continuous M M'" "AE x in M. P x"
    shows "AE x in M'. P x"
 proof -
-  interpret \<nu>: measure_space M' 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
+  from `AE x in M. P x` obtain N where N: "N \<in> null_sets M" "{x\<in>space M. \<not> P x} \<subseteq> N"
+    unfolding eventually_ae_filter by auto
   show "AE x in M'. P x"
-  proof (rule \<nu>.AE_I')
-    show "{x\<in>space M'. \<not> P x} \<subseteq> N" by simp fact
-    from `absolutely_continuous (measure M')` show "N \<in> \<nu>.null_sets"
-      using N unfolding absolutely_continuous_def by auto
+  proof (rule AE_I')
+    show "{x\<in>space M'. \<not> P x} \<subseteq> N" using sets_eq_imp_space_eq[OF sets_eq] N(2) by simp
+    from `absolutely_continuous M M'` show "N \<in> null_sets M'"
+      using N unfolding absolutely_continuous_def sets_eq null_sets_def by auto
   qed
 qed
 
-lemma (in finite_measure_space) absolutely_continuousI:
-  assumes "finite_measure_space (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure_space ?\<nu>")
-  assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
-  shows "absolutely_continuous \<nu>"
-proof (unfold absolutely_continuous_def sets_eq_Pow, safe)
-  fix N assume "\<mu> N = 0" "N \<subseteq> space M"
-  interpret v: finite_measure_space ?\<nu> by fact
-  have "\<nu> N = measure ?\<nu> (\<Union>x\<in>N. {x})" by simp
-  also have "\<dots> = (\<Sum>x\<in>N. measure ?\<nu> {x})"
-  proof (rule v.measure_setsum[symmetric])
-    show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset)
-    show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto
-    fix x assume "x \<in> N" thus "{x} \<in> sets ?\<nu>" using `N \<subseteq> space M` sets_eq_Pow by auto
-  qed
-  also have "\<dots> = 0"
-  proof (safe intro!: setsum_0')
-    fix x assume "x \<in> N"
-    hence "\<mu> {x} \<le> \<mu> N" "0 \<le> \<mu> {x}"
-      using sets_eq_Pow `N \<subseteq> space M` positive_measure[of "{x}"]
-      by (auto intro!: measure_mono)
-    then have "\<mu> {x} = 0" using `\<mu> N = 0` by simp
-    thus "measure ?\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
-  qed
-  finally show "\<nu> N = 0" by simp
-qed
-
-lemma (in measure_space) density_is_absolutely_continuous:
-  assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
-  shows "absolutely_continuous \<nu>"
-  using assms unfolding absolutely_continuous_def
-  by (simp add: positive_integral_null_set)
-
 subsection "Existence of the Radon-Nikodym derivative"
 
 lemma (in finite_measure) Radon_Nikodym_aux_epsilon:
   fixes e :: real assumes "0 < e"
-  assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)"
-  shows "\<exists>A\<in>sets M. \<mu>' (space M) - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) (space M) \<le>
-                    \<mu>' A - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) A \<and>
-                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < \<mu>' B - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) B)"
+  assumes "finite_measure N" and sets_eq: "sets N = sets M"
+  shows "\<exists>A\<in>sets M. measure M (space M) - measure N (space M) \<le> measure M A - measure N A \<and>
+                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> - e < measure M B - measure N B)"
 proof -
-  interpret M': finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
-  let ?d = "\<lambda>A. \<mu>' A - M'.\<mu>' A"
+  interpret M': finite_measure N by fact
+  let ?d = "\<lambda>A. measure M A - measure N A"
   let ?A = "\<lambda>A. if (\<forall>B\<in>sets M. B \<subseteq> space M - A \<longrightarrow> -e < ?d B)
     then {}
     else (SOME B. B \<in> sets M \<and> B \<subseteq> space M - A \<and> ?d B \<le> -e)"
@@ -159,7 +170,7 @@
       fix B assume "B \<in> sets M \<and> B \<subseteq> space M - A n \<and> ?d B \<le> - e"
       hence "A n \<inter> B = {}" "B \<in> sets M" and dB: "?d B \<le> -e" by auto
       hence "?d (A n \<union> B) = ?d (A n) + ?d B"
-        using `A n \<in> sets M` finite_measure_Union M'.finite_measure_Union by simp
+        using `A n \<in> sets M` finite_measure_Union M'.finite_measure_Union by (simp add: sets_eq)
       also have "\<dots> \<le> ?d (A n) - e" using dB by simp
       finally show "?d (A n \<union> B) \<le> ?d (A n) - e" .
     qed }
@@ -187,8 +198,8 @@
       proof (induct n)
         fix n assume "?d (space M) \<le> ?d (space M - A n)"
         also have "\<dots> \<le> ?d (space M - A (Suc n))"
-          using A_in_sets sets_into_space dA_mono[of n]
-          by (simp del: A_simps add: finite_measure_Diff M'.finite_measure_Diff)
+          using A_in_sets sets_into_space dA_mono[of n] finite_measure_compl M'.finite_measure_compl
+          by (simp del: A_simps add: sets_eq sets_eq_imp_space_eq[OF sets_eq])
         finally show "?d (space M) \<le> ?d (space M - A (Suc n))" .
       qed simp
     qed
@@ -199,17 +210,17 @@
       proof (induct n)
         case (Suc n) with dA_epsilon[of n, OF B] show ?case by (simp del: A_simps add: real_of_nat_Suc field_simps)
       next
-        case 0 with M'.empty_measure show ?case by (simp add: zero_ereal_def)
+        case 0 with measure_empty show ?case by (simp add: zero_ereal_def)
       qed } note dA_less = this
     have decseq: "decseq (\<lambda>n. ?d (A n))" unfolding decseq_eq_incseq
     proof (rule incseq_SucI)
       fix n show "- ?d (A n) \<le> - ?d (A (Suc n))" using dA_mono[of n] by auto
     qed
     have A: "incseq A" by (auto intro!: incseq_SucI)
-    from finite_continuity_from_below[OF _ A] `range A \<subseteq> sets M`
-      M'.finite_continuity_from_below[OF _ A]
+    from finite_Lim_measure_incseq[OF _ A] `range A \<subseteq> sets M`
+      M'.finite_Lim_measure_incseq[OF _ A]
     have convergent: "(\<lambda>i. ?d (A i)) ----> ?d (\<Union>i. A i)"
-      by (auto intro!: tendsto_diff)
+      by (auto intro!: tendsto_diff simp: sets_eq)
     obtain n :: nat where "- ?d (\<Union>i. A i) / e < real n" using reals_Archimedean2 by auto
     moreover from order_trans[OF decseq_le[OF decseq convergent] dA_less]
     have "real n \<le> - ?d (\<Union>i. A i) / e" using `0<e` by (simp add: field_simps)
@@ -217,57 +228,24 @@
   qed
 qed
 
-lemma (in finite_measure) restricted_measure_subset:
-  assumes S: "S \<in> sets M" and X: "X \<subseteq> S"
-  shows "finite_measure.\<mu>' (restricted_space S) X = \<mu>' X"
-proof cases
-  note r = restricted_finite_measure[OF S]
-  { assume "X \<in> sets M" with S X show ?thesis
-      unfolding finite_measure.\<mu>'_def[OF r] \<mu>'_def by auto }
-  { assume "X \<notin> sets M"
-    moreover then have "S \<inter> X \<notin> sets M"
-      using X by (simp add: Int_absorb1)
-    ultimately show ?thesis
-      unfolding finite_measure.\<mu>'_def[OF r] \<mu>'_def using S by auto }
-qed
-
-lemma (in finite_measure) restricted_measure:
-  assumes X: "S \<in> sets M" "X \<in> sets (restricted_space S)"
-  shows "finite_measure.\<mu>' (restricted_space S) X = \<mu>' X"
+lemma (in finite_measure) Radon_Nikodym_aux:
+  assumes "finite_measure N" and sets_eq: "sets N = sets M"
+  shows "\<exists>A\<in>sets M. measure M (space M) - measure N (space M) \<le>
+                    measure M A - measure N A \<and>
+                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> measure M B - measure N B)"
 proof -
-  from X have "S \<in> sets M" "X \<subseteq> S" by auto
-  from restricted_measure_subset[OF this] show ?thesis .
-qed
-
-lemma (in finite_measure) Radon_Nikodym_aux:
-  assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?M'")
-  shows "\<exists>A\<in>sets M. \<mu>' (space M) - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) (space M) \<le>
-                    \<mu>' A - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) A \<and>
-                    (\<forall>B\<in>sets M. B \<subseteq> A \<longrightarrow> 0 \<le> \<mu>' B - finite_measure.\<mu>' (M\<lparr>measure := \<nu>\<rparr>) B)"
-proof -
-  interpret M': finite_measure ?M' where
-    "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>" by fact auto
-  let ?d = "\<lambda>A. \<mu>' A - M'.\<mu>' A"
+  interpret N: finite_measure N by fact
+  let ?d = "\<lambda>A. measure M A - measure N A"
   let ?P = "\<lambda>A B n. A \<in> sets M \<and> A \<subseteq> B \<and> ?d B \<le> ?d A \<and> (\<forall>C\<in>sets M. C \<subseteq> A \<longrightarrow> - 1 / real (Suc n) < ?d C)"
   let ?r = "\<lambda>S. restricted_space S"
   { fix S n assume S: "S \<in> sets M"
-    note r = M'.restricted_finite_measure[of S] restricted_finite_measure[OF S] S
-    then have "finite_measure (?r S)" "0 < 1 / real (Suc n)"
-      "finite_measure (?r S\<lparr>measure := \<nu>\<rparr>)" by auto
+    then have "finite_measure (density M (indicator S))" "0 < 1 / real (Suc n)"
+         "finite_measure (density N (indicator S))" "sets (density N (indicator S)) = sets (density M (indicator S))"
+      by (auto simp: finite_measure_restricted N.finite_measure_restricted sets_eq)
     from finite_measure.Radon_Nikodym_aux_epsilon[OF this] guess X .. note X = this
-    have "?P X S n"
-    proof (intro conjI ballI impI)
-      show "X \<in> sets M" "X \<subseteq> S" using X(1) `S \<in> sets M` by auto
-      have "S \<in> op \<inter> S ` sets M" using `S \<in> sets M` by auto
-      then show "?d S \<le> ?d X"
-        using S X restricted_measure[OF S] M'.restricted_measure[OF S] by simp
-      fix C assume "C \<in> sets M" "C \<subseteq> X"
-      then have "C \<in> sets (restricted_space S)" "C \<subseteq> X"
-        using `S \<in> sets M` `X \<subseteq> S` by auto
-      with X(2) show "- 1 / real (Suc n) < ?d C"
-        using S X restricted_measure[OF S] M'.restricted_measure[OF S] by auto
-    qed
-    hence "\<exists>A. ?P A S n" by auto }
+    with S have "?P (S \<inter> X) S n"
+      by (simp add: measure_restricted sets_eq Int) (metis inf_absorb2)
+    hence "\<exists>A. ?P A S n" .. }
   note Ex_P = this
   def A \<equiv> "nat_rec (space M) (\<lambda>n A. SOME B. ?P B A n)"
   have A_Suc: "\<And>n. A (Suc n) = (SOME B. ?P B (A n) n)" by (simp add: A_def)
@@ -292,12 +270,11 @@
   proof (safe intro!: bexI[of _ "\<Inter>i. A i"])
     show "(\<Inter>i. A i) \<in> sets M" using A_in_sets by auto
     have A: "decseq A" using A_mono by (auto intro!: decseq_SucI)
-    from
-      finite_continuity_from_above[OF `range A \<subseteq> sets M` A]
-      M'.finite_continuity_from_above[OF `range A \<subseteq> sets M` A]
-    have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (intro tendsto_diff)
+    from `range A \<subseteq> sets M`
+      finite_Lim_measure_decseq[OF _ A] N.finite_Lim_measure_decseq[OF _ A]
+    have "(\<lambda>i. ?d (A i)) ----> ?d (\<Inter>i. A i)" by (auto intro!: tendsto_diff simp: sets_eq)
     thus "?d (space M) \<le> ?d (\<Inter>i. A i)" using mono_dA[THEN monoD, of 0 _]
-      by (rule_tac LIMSEQ_le_const) (auto intro!: exI)
+      by (rule_tac LIMSEQ_le_const) auto
   next
     fix B assume B: "B \<in> sets M" "B \<subseteq> (\<Inter>i. A i)"
     show "0 \<le> ?d B"
@@ -315,14 +292,12 @@
 qed
 
 lemma (in finite_measure) Radon_Nikodym_finite_measure:
-  assumes "finite_measure (M\<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?M'")
-  assumes "absolutely_continuous \<nu>"
-  shows "\<exists>f \<in> borel_measurable M. \<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
+  assumes "finite_measure N" and sets_eq: "sets N = sets M"
+  assumes "absolutely_continuous M N"
+  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"
 proof -
-  interpret M': finite_measure ?M'
-    where "space ?M' = space M" and "sets ?M' = sets M" and "measure ?M' = \<nu>"
-    using assms(1) by auto
-  def G \<equiv> "{g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A)}"
+  interpret N: finite_measure N by fact
+  def G \<equiv> "{g \<in> borel_measurable M. (\<forall>x. 0 \<le> g x) \<and> (\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> N A)}"
   have "(\<lambda>x. 0) \<in> G" unfolding G_def by auto
   hence "G \<noteq> {}" by auto
   { fix f g assume f: "f \<in> G" and g: "g \<in> G"
@@ -333,6 +308,7 @@
       have "?A \<in> sets M" using f g unfolding G_def by auto
       fix A assume "A \<in> sets M"
       hence sets: "?A \<inter> A \<in> sets M" "(space M - ?A) \<inter> A \<in> sets M" using `?A \<in> sets M` by auto
+      hence sets': "?A \<inter> A \<in> sets N" "(space M - ?A) \<inter> A \<in> sets N" by (auto simp: sets_eq)
       have union: "((?A \<inter> A) \<union> ((space M - ?A) \<inter> A)) = A"
         using sets_into_space[OF `A \<in> sets M`] by auto
       have "\<And>x. x \<in> space M \<Longrightarrow> max (g x) (f x) * indicator A x =
@@ -343,11 +319,11 @@
         (\<integral>\<^isup>+x. f x * indicator ((space M - ?A) \<inter> A) x \<partial>M)"
         using f g sets unfolding G_def
         by (auto cong: positive_integral_cong intro!: positive_integral_add)
-      also have "\<dots> \<le> \<nu> (?A \<inter> A) + \<nu> ((space M - ?A) \<inter> A)"
+      also have "\<dots> \<le> N (?A \<inter> A) + N ((space M - ?A) \<inter> A)"
         using f g sets unfolding G_def by (auto intro!: add_mono)
-      also have "\<dots> = \<nu> A"
-        using M'.measure_additive[OF sets] union by auto
-      finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> \<nu> A" .
+      also have "\<dots> = N A"
+        using plus_emeasure[OF sets'] union by auto
+      finally show "(\<integral>\<^isup>+x. max (g x) (f x) * indicator A x \<partial>M) \<le> N A" .
     next
       fix x show "0 \<le> max (g x) (f x)" using f g by (auto simp: G_def split: split_max)
     qed }
@@ -368,15 +344,15 @@
         using `incseq f` f `A \<in> sets M`
         by (intro positive_integral_monotone_convergence_SUP)
            (auto simp: G_def incseq_Suc_iff le_fun_def split: split_indicator)
-      finally show "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> \<nu> A"
+      finally show "(\<integral>\<^isup>+x. (SUP i. f i x) * indicator A x \<partial>M) \<le> N A"
         using f `A \<in> sets M` by (auto intro!: SUP_least simp: G_def)
     qed }
   note SUP_in_G = this
   let ?y = "SUP g : G. integral\<^isup>P M g"
-  have "?y \<le> \<nu> (space M)" unfolding G_def
+  have y_le: "?y \<le> N (space M)" unfolding G_def
   proof (safe intro!: SUP_least)
-    fix g assume "\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> \<nu> A"
-    from this[THEN bspec, OF top] show "integral\<^isup>P M g \<le> \<nu> (space M)"
+    fix g assume "\<forall>A\<in>sets M. (\<integral>\<^isup>+x. g x * indicator A x \<partial>M) \<le> N A"
+    from this[THEN bspec, OF top] show "integral\<^isup>P M g \<le> N (space M)"
       by (simp cong: positive_integral_cong)
   qed
   from SUPR_countable_SUPR[OF `G \<noteq> {}`, of "integral\<^isup>P M"] guess ys .. note ys = this
@@ -411,8 +387,7 @@
   also have "\<dots> = ?y"
   proof (rule antisym)
     show "(SUP i. integral\<^isup>P M (?g i)) \<le> ?y"
-      using g_in_G
-      by (auto intro!: exI Sup_mono simp: SUP_def)
+      using g_in_G by (auto intro: Sup_mono simp: SUP_def)
     show "?y \<le> (SUP i. integral\<^isup>P M (?g i))" unfolding y_eq
       by (auto intro!: SUP_mono positive_integral_mono Max_ge)
   qed
@@ -420,107 +395,68 @@
   have "\<And>x. 0 \<le> f x"
     unfolding f_def using `\<And>i. gs i \<in> G`
     by (auto intro!: SUP_upper2 Max_ge_iff[THEN iffD2] simp: G_def)
-  let ?t = "\<lambda>A. \<nu> A - (\<integral>\<^isup>+x. ?F A x \<partial>M)"
-  let ?M = "M\<lparr> measure := ?t\<rparr>"
-  interpret M: sigma_algebra ?M
-    by (intro sigma_algebra_cong) auto
-  have f_le_\<nu>: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> \<nu> A"
+  let ?t = "\<lambda>A. N A - (\<integral>\<^isup>+x. ?F A x \<partial>M)"
+  let ?M = "diff_measure N (density M f)"
+  have f_le_N: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> N A"
     using `f \<in> G` unfolding G_def by auto
-  have fmM: "finite_measure ?M"
+  have emeasure_M: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure ?M A = ?t A"
+  proof (subst emeasure_diff_measure)
+    from f_le_N[of "space M"] show "finite_measure N" "finite_measure (density M f)"
+      by (auto intro!: finite_measureI simp: emeasure_density cong: positive_integral_cong)
+  next
+    fix B assume "B \<in> sets N" with f_le_N[of B] show "emeasure (density M f) B \<le> emeasure N B"
+      by (auto simp: sets_eq emeasure_density cong: positive_integral_cong)
+  qed (auto simp: sets_eq emeasure_density)
+  from emeasure_M[of "space M"] N.finite_emeasure_space positive_integral_positive[of M "?F (space M)"]
+  interpret M': finite_measure ?M
+    by (auto intro!: finite_measureI simp: sets_eq_imp_space_eq[OF sets_eq] N.emeasure_eq_measure ereal_minus_eq_PInfty_iff)
+
+  have ac: "absolutely_continuous M ?M" unfolding absolutely_continuous_def
   proof
-    show "measure_space ?M"
-    proof (default, simp_all add: countably_additive_def positive_def, safe del: notI)
-      fix A :: "nat \<Rightarrow> 'a set"  assume A: "range A \<subseteq> sets M" "disjoint_family A"
-      have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. (\<Sum>n. ?F (A n) x) \<partial>M)"
-        using `range A \<subseteq> sets M` `\<And>x. 0 \<le> f x`
-        by (intro positive_integral_suminf[symmetric]) auto
-      also have "\<dots> = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)"
-        using `\<And>x. 0 \<le> f x`
-        by (intro positive_integral_cong) (simp add: suminf_cmult_ereal suminf_indicator[OF `disjoint_family A`])
-      finally have "(\<Sum>n. (\<integral>\<^isup>+x. ?F (A n) x \<partial>M)) = (\<integral>\<^isup>+x. ?F (\<Union>n. A n) x \<partial>M)" .
-      moreover have "(\<Sum>n. \<nu> (A n)) = \<nu> (\<Union>n. A n)"
-        using M'.measure_countably_additive A by (simp add: comp_def)
-      moreover have v_fin: "\<nu> (\<Union>i. A i) \<noteq> \<infinity>" using M'.finite_measure A by (simp add: countable_UN)
-      moreover {
-        have "(\<integral>\<^isup>+x. ?F (\<Union>i. A i) x \<partial>M) \<le> \<nu> (\<Union>i. A i)"
-          using A `f \<in> G` unfolding G_def by (auto simp: countable_UN)
-        also have "\<nu> (\<Union>i. A i) < \<infinity>" using v_fin by simp
-        finally have "(\<integral>\<^isup>+x. ?F (\<Union>i. A i) x \<partial>M) \<noteq> \<infinity>" by simp }
-      moreover have "\<And>i. (\<integral>\<^isup>+x. ?F (A i) x \<partial>M) \<le> \<nu> (A i)"
-        using A by (intro f_le_\<nu>) auto
-      ultimately
-      show "(\<Sum>n. ?t (A n)) = ?t (\<Union>i. A i)"
-        by (subst suminf_ereal_minus) (simp_all add: positive_integral_positive)
-    next
-      fix A assume A: "A \<in> sets M" show "0 \<le> \<nu> A - \<integral>\<^isup>+ x. ?F A x \<partial>M"
-        using f_le_\<nu>[OF A] `f \<in> G` M'.finite_measure[OF A] by (auto simp: G_def ereal_le_minus_iff)
-    qed
-  next
-    show "measure ?M (space ?M) \<noteq> \<infinity>"
-      using positive_integral_positive[of "?F (space M)"]
-      by (cases rule: ereal2_cases[of "\<nu> (space M)" "\<integral>\<^isup>+ x. ?F (space M) x \<partial>M"]) auto
+    fix A assume A: "A \<in> null_sets M"
+    with `absolutely_continuous M N` have "A \<in> null_sets N"
+      unfolding absolutely_continuous_def by auto
+    moreover with A have "(\<integral>\<^isup>+ x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` by (auto simp: G_def)
+    ultimately have "N A - (\<integral>\<^isup>+ x. ?F A x \<partial>M) = 0"
+      using positive_integral_positive[of M] by (auto intro!: antisym)
+    then show "A \<in> null_sets ?M"
+      using A by (simp add: emeasure_M null_sets_def sets_eq)
   qed
-  then interpret M: finite_measure ?M
-    where "space ?M = space M" and "sets ?M = sets M" and "measure ?M = ?t"
-    by (simp_all add: fmM)
-  have ac: "absolutely_continuous ?t" unfolding absolutely_continuous_def
-  proof
-    fix N assume N: "N \<in> null_sets"
-    with `absolutely_continuous \<nu>` have "\<nu> N = 0" unfolding absolutely_continuous_def by auto
-    moreover with N have "(\<integral>\<^isup>+ x. ?F N x \<partial>M) \<le> \<nu> N" using `f \<in> G` by (auto simp: G_def)
-    ultimately show "\<nu> N - (\<integral>\<^isup>+ x. ?F N x \<partial>M) = 0"
-      using positive_integral_positive by (auto intro!: antisym)
-  qed
-  have upper_bound: "\<forall>A\<in>sets M. ?t A \<le> 0"
+  have upper_bound: "\<forall>A\<in>sets M. ?M A \<le> 0"
   proof (rule ccontr)
     assume "\<not> ?thesis"
-    then obtain A where A: "A \<in> sets M" and pos: "0 < ?t A"
+    then obtain A where A: "A \<in> sets M" and pos: "0 < ?M A"
       by (auto simp: not_le)
     note pos
-    also have "?t A \<le> ?t (space M)"
-      using M.measure_mono[of A "space M"] A sets_into_space by simp
-    finally have pos_t: "0 < ?t (space M)" by simp
+    also have "?M A \<le> ?M (space M)"
+      using emeasure_space[of ?M A] by (simp add: sets_eq[THEN sets_eq_imp_space_eq])
+    finally have pos_t: "0 < ?M (space M)" by simp
     moreover
-    then have "\<mu> (space M) \<noteq> 0"
-      using ac unfolding absolutely_continuous_def by auto
-    then have pos_M: "0 < \<mu> (space M)"
-      using positive_measure[OF top] by (simp add: le_less)
+    then have "emeasure M (space M) \<noteq> 0"
+      using ac unfolding absolutely_continuous_def by (auto simp: null_sets_def)
+    then have pos_M: "0 < emeasure M (space M)"
+      using emeasure_nonneg[of M "space M"] by (simp add: le_less)
     moreover
-    have "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<le> \<nu> (space M)"
+    have "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<le> N (space M)"
       using `f \<in> G` unfolding G_def by auto
     hence "(\<integral>\<^isup>+x. f x * indicator (space M) x \<partial>M) \<noteq> \<infinity>"
-      using M'.finite_measure_of_space by auto
+      using M'.finite_emeasure_space by auto
     moreover
-    def b \<equiv> "?t (space M) / \<mu> (space M) / 2"
+    def b \<equiv> "?M (space M) / emeasure M (space M) / 2"
     ultimately have b: "b \<noteq> 0 \<and> 0 \<le> b \<and> b \<noteq> \<infinity>"
-      using M'.finite_measure_of_space positive_integral_positive[of "?F (space M)"]
-      by (cases rule: ereal3_cases[of "integral\<^isup>P M (?F (space M))" "\<nu> (space M)" "\<mu> (space M)"])
-         (simp_all add: field_simps)
+      by (auto simp: ereal_divide_eq)
     then have b: "b \<noteq> 0" "0 \<le> b" "0 < b"  "b \<noteq> \<infinity>" by auto
-    let ?Mb = "?M\<lparr>measure := \<lambda>A. b * \<mu> A\<rparr>"
-    interpret b: sigma_algebra ?Mb by (intro sigma_algebra_cong) auto
-    have Mb: "finite_measure ?Mb"
-    proof
-      show "measure_space ?Mb"
-      proof
-        show "positive ?Mb (measure ?Mb)"
-          using `0 \<le> b` by (auto simp: positive_def)
-        show "countably_additive ?Mb (measure ?Mb)"
-          using `0 \<le> b` measure_countably_additive
-          by (auto simp: countably_additive_def suminf_cmult_ereal subset_eq)
-      qed
-      show "measure ?Mb (space ?Mb) \<noteq> \<infinity>"
-        using b by auto
-    qed
-    from M.Radon_Nikodym_aux[OF this]
-    obtain A0 where "A0 \<in> sets M" and
-      space_less_A0: "real (?t (space M)) - real (b * \<mu> (space M)) \<le> real (?t A0) - real (b * \<mu> A0)" and
-      *: "\<And>B. \<lbrakk> B \<in> sets M ; B \<subseteq> A0 \<rbrakk> \<Longrightarrow> 0 \<le> real (?t B) - real (b * \<mu> B)"
-      unfolding M.\<mu>'_def finite_measure.\<mu>'_def[OF Mb] by auto
+    let ?Mb = "density M (\<lambda>_. b)"
+    have Mb: "finite_measure ?Mb" "sets ?Mb = sets ?M"
+        using b by (auto simp: emeasure_density_const sets_eq intro!: finite_measureI)
+    from M'.Radon_Nikodym_aux[OF this] guess A0 ..
+    then have "A0 \<in> sets M"
+      and space_less_A0: "measure ?M (space M) - real b * measure M (space M) \<le> measure ?M A0 - real b * measure M A0"
+      and *: "\<And>B. B \<in> sets M \<Longrightarrow> B \<subseteq> A0 \<Longrightarrow> 0 \<le> measure ?M B - real b * measure M B"
+      using b by (simp_all add: measure_density_const sets_eq_imp_space_eq[OF sets_eq] sets_eq)
     { fix B assume B: "B \<in> sets M" "B \<subseteq> A0"
-      with *[OF this] have "b * \<mu> B \<le> ?t B"
-        using M'.finite_measure b finite_measure M.positive_measure[OF B(1)]
-        by (cases rule: ereal2_cases[of "?t B" "b * \<mu> B"]) auto }
+      with *[OF this] have "b * emeasure M B \<le> ?M B"
+        using b unfolding M'.emeasure_eq_measure emeasure_eq_measure by (cases b) auto }
     note bM_le_t = this
     let ?f0 = "\<lambda>x. f x + b * indicator A0 x"
     { fix A assume A: "A \<in> sets M"
@@ -529,71 +465,47 @@
         (\<integral>\<^isup>+x. f x * indicator A x + b * indicator (A \<inter> A0) x \<partial>M)"
         by (auto intro!: positive_integral_cong split: split_indicator)
       hence "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) =
-          (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0)"
+          (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * emeasure M (A \<inter> A0)"
         using `A0 \<in> sets M` `A \<inter> A0 \<in> sets M` A b `f \<in> G`
         by (simp add: G_def positive_integral_add positive_integral_cmult_indicator) }
     note f0_eq = this
     { fix A assume A: "A \<in> sets M"
       hence "A \<inter> A0 \<in> sets M" using `A0 \<in> sets M` by auto
-      have f_le_v: "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A"
-        using `f \<in> G` A unfolding G_def by auto
+      have f_le_v: "(\<integral>\<^isup>+x. ?F A x \<partial>M) \<le> N A" using `f \<in> G` A unfolding G_def by auto
       note f0_eq[OF A]
-      also have "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + b * \<mu> (A \<inter> A0) \<le>
-          (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t (A \<inter> A0)"
+      also have "(\<integral>\<^isup>+x. ?F A x \<partial>M) + b * emeasure M (A \<inter> A0) \<le> (\<integral>\<^isup>+x. ?F A x \<partial>M) + ?M (A \<inter> A0)"
         using bM_le_t[OF `A \<inter> A0 \<in> sets M`] `A \<in> sets M` `A0 \<in> sets M`
         by (auto intro!: add_left_mono)
-      also have "\<dots> \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?t A"
-        using M.measure_mono[simplified, OF _ `A \<inter> A0 \<in> sets M` `A \<in> sets M`]
-        by (auto intro!: add_left_mono)
-      also have "\<dots> \<le> \<nu> A"
-        using f_le_v M'.finite_measure[simplified, OF `A \<in> sets M`] positive_integral_positive[of "?F A"]
-        by (cases "\<integral>\<^isup>+x. ?F A x \<partial>M", cases "\<nu> A") auto
-      finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> \<nu> A" . }
+      also have "\<dots> \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) + ?M A"
+        using emeasure_mono[of "A \<inter> A0" A ?M] `A \<in> sets M` `A0 \<in> sets M`
+        by (auto intro!: add_left_mono simp: sets_eq)
+      also have "\<dots> \<le> N A"
+        unfolding emeasure_M[OF `A \<in> sets M`]
+        using f_le_v N.emeasure_eq_measure[of A] positive_integral_positive[of M "?F A"]
+        by (cases "\<integral>\<^isup>+x. ?F A x \<partial>M", cases "N A") auto
+      finally have "(\<integral>\<^isup>+x. ?f0 x * indicator A x \<partial>M) \<le> N A" . }
     hence "?f0 \<in> G" using `A0 \<in> sets M` b `f \<in> G` unfolding G_def
       by (auto intro!: ereal_add_nonneg_nonneg)
-    have real: "?t (space M) \<noteq> \<infinity>" "?t A0 \<noteq> \<infinity>"
-      "b * \<mu> (space M) \<noteq> \<infinity>" "b * \<mu> A0 \<noteq> \<infinity>"
-      using `A0 \<in> sets M` b
-        finite_measure[of A0] M.finite_measure[of A0]
-        finite_measure_of_space M.finite_measure_of_space
-      by auto
     have int_f_finite: "integral\<^isup>P M f \<noteq> \<infinity>"
-      using M'.finite_measure_of_space pos_t unfolding ereal_less_minus_iff
-      by (auto cong: positive_integral_cong)
-    have  "0 < ?t (space M) - b * \<mu> (space M)" unfolding b_def
-      using finite_measure_of_space M'.finite_measure_of_space pos_t pos_M
-      using positive_integral_positive[of "?F (space M)"]
-      by (cases rule: ereal3_cases[of "\<mu> (space M)" "\<nu> (space M)" "integral\<^isup>P M (?F (space M))"])
-         (auto simp: field_simps mult_less_cancel_left)
-    also have "\<dots> \<le> ?t A0 - b * \<mu> A0"
-      using space_less_A0 b
-      using
-        `A0 \<in> sets M`[THEN M.real_measure]
-        top[THEN M.real_measure]
-      apply safe
-      apply simp
-      using
-        `A0 \<in> sets M`[THEN real_measure]
-        `A0 \<in> sets M`[THEN M'.real_measure]
-        top[THEN real_measure]
-        top[THEN M'.real_measure]
-      by (cases b) auto
-    finally have 1: "b * \<mu> A0 < ?t A0"
-      using
-        `A0 \<in> sets M`[THEN M.real_measure]
-      apply safe
-      apply simp
-      using
-        `A0 \<in> sets M`[THEN real_measure]
-        `A0 \<in> sets M`[THEN M'.real_measure]
-      by (cases b) auto
-    have "0 < ?t A0"
-      using b `A0 \<in> sets M` by (auto intro!: le_less_trans[OF _ 1])
-    then have "\<mu> A0 \<noteq> 0" using ac unfolding absolutely_continuous_def
-      using `A0 \<in> sets M` by auto
-    then have "0 < \<mu> A0" using positive_measure[OF `A0 \<in> sets M`] by auto
-    hence "0 < b * \<mu> A0" using b by (auto simp: ereal_zero_less_0_iff)
-    with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * \<mu> A0" unfolding int_f_eq_y
+      by (metis N.emeasure_finite ereal_infty_less_eq2(1) int_f_eq_y y_le)
+    have  "0 < ?M (space M) - emeasure ?Mb (space M)"
+      using pos_t
+      by (simp add: b emeasure_density_const)
+         (simp add: M'.emeasure_eq_measure emeasure_eq_measure pos_M b_def)
+    also have "\<dots> \<le> ?M A0 - b * emeasure M A0"
+      using space_less_A0 `A0 \<in> sets M` b
+      by (cases b) (auto simp add: b emeasure_density_const sets_eq M'.emeasure_eq_measure emeasure_eq_measure)
+    finally have 1: "b * emeasure M A0 < ?M A0"
+      by (metis M'.emeasure_real `A0 \<in> sets M` bM_le_t diff_self ereal_less(1) ereal_minus(1)
+                less_eq_ereal_def mult_zero_left not_square_less_zero subset_refl zero_ereal_def)
+    with b have "0 < ?M A0"
+      by (metis M'.emeasure_real MInfty_neq_PInfty(1) emeasure_real ereal_less_eq(5) ereal_zero_times
+               ereal_mult_eq_MInfty ereal_mult_eq_PInfty ereal_zero_less_0_iff less_eq_ereal_def)
+    then have "emeasure M A0 \<noteq> 0" using ac `A0 \<in> sets M`
+      by (auto simp: absolutely_continuous_def null_sets_def)
+    then have "0 < emeasure M A0" using emeasure_nonneg[of M A0] by auto
+    hence "0 < b * emeasure M A0" using b by (auto simp: ereal_zero_less_0_iff)
+    with int_f_finite have "?y + 0 < integral\<^isup>P M f + b * emeasure M A0" unfolding int_f_eq_y
       using `f \<in> G`
       by (intro ereal_add_strict_mono) (auto intro!: SUP_upper2 positive_integral_positive)
     also have "\<dots> = integral\<^isup>P M ?f0" using f0_eq[OF top] `A0 \<in> sets M` sets_into_space
@@ -602,75 +514,67 @@
     moreover from `?f0 \<in> G` have "integral\<^isup>P M ?f0 \<le> ?y" by (auto intro!: SUP_upper)
     ultimately show False by auto
   qed
+  let ?f = "\<lambda>x. max 0 (f x)"
   show ?thesis
-  proof (safe intro!: bexI[of _ f])
-    fix A assume A: "A\<in>sets M"
-    show "\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
-    proof (rule antisym)
-      show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) \<le> \<nu> A"
-        using `f \<in> G` `A \<in> sets M` unfolding G_def by auto
-      show "\<nu> A \<le> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
-        using upper_bound[THEN bspec, OF `A \<in> sets M`]
-        using M'.real_measure[OF A]
-        by (cases "integral\<^isup>P M (?F A)") auto
-    qed
-  qed simp
+  proof (intro bexI[of _ ?f] measure_eqI conjI)
+    show "sets (density M ?f) = sets N"
+      by (simp add: sets_eq)
+    fix A assume A: "A\<in>sets (density M ?f)"
+    then show "emeasure (density M ?f) A = emeasure N A"
+      using `f \<in> G` A upper_bound[THEN bspec, of A] N.emeasure_eq_measure[of A]
+      by (cases "integral\<^isup>P M (?F A)")
+         (auto intro!: antisym simp add: emeasure_density G_def emeasure_M density_max_0[symmetric])
+  qed auto
 qed
 
 lemma (in finite_measure) split_space_into_finite_sets_and_rest:
-  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
-  assumes ac: "absolutely_continuous \<nu>"
+  assumes ac: "absolutely_continuous M N" and sets_eq: "sets N = sets M"
   shows "\<exists>A0\<in>sets M. \<exists>B::nat\<Rightarrow>'a set. disjoint_family B \<and> range B \<subseteq> sets M \<and> A0 = space M - (\<Union>i. B i) \<and>
-    (\<forall>A\<in>sets M. A \<subseteq> A0 \<longrightarrow> (\<mu> A = 0 \<and> \<nu> A = 0) \<or> (\<mu> A > 0 \<and> \<nu> A = \<infinity>)) \<and>
-    (\<forall>i. \<nu> (B i) \<noteq> \<infinity>)"
+    (\<forall>A\<in>sets M. A \<subseteq> A0 \<longrightarrow> (emeasure M A = 0 \<and> N A = 0) \<or> (emeasure M A > 0 \<and> N A = \<infinity>)) \<and>
+    (\<forall>i. N (B i) \<noteq> \<infinity>)"
 proof -
-  interpret v: measure_space ?N
-    where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>"
-    by fact auto
-  let ?Q = "{Q\<in>sets M. \<nu> Q \<noteq> \<infinity>}"
-  let ?a = "SUP Q:?Q. \<mu> Q"
-  have "{} \<in> ?Q" using v.empty_measure by auto
+  let ?Q = "{Q\<in>sets M. N Q \<noteq> \<infinity>}"
+  let ?a = "SUP Q:?Q. emeasure M Q"
+  have "{} \<in> ?Q" by auto
   then have Q_not_empty: "?Q \<noteq> {}" by blast
-  have "?a \<le> \<mu> (space M)" using sets_into_space
-    by (auto intro!: SUP_least measure_mono)
-  then have "?a \<noteq> \<infinity>" using finite_measure_of_space
+  have "?a \<le> emeasure M (space M)" using sets_into_space
+    by (auto intro!: SUP_least emeasure_mono)
+  then have "?a \<noteq> \<infinity>" using finite_emeasure_space
     by auto
-  from SUPR_countable_SUPR[OF Q_not_empty, of \<mu>]
-  obtain Q'' where "range Q'' \<subseteq> \<mu> ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
+  from SUPR_countable_SUPR[OF Q_not_empty, of "emeasure M"]
+  obtain Q'' where "range Q'' \<subseteq> emeasure M ` ?Q" and a: "?a = (SUP i::nat. Q'' i)"
     by auto
-  then have "\<forall>i. \<exists>Q'. Q'' i = \<mu> Q' \<and> Q' \<in> ?Q" by auto
-  from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = \<mu> (Q' i)" "\<And>i. Q' i \<in> ?Q"
+  then have "\<forall>i. \<exists>Q'. Q'' i = emeasure M Q' \<and> Q' \<in> ?Q" by auto
+  from choice[OF this] obtain Q' where Q': "\<And>i. Q'' i = emeasure M (Q' i)" "\<And>i. Q' i \<in> ?Q"
     by auto
-  then have a_Lim: "?a = (SUP i::nat. \<mu> (Q' i))" using a by simp
+  then have a_Lim: "?a = (SUP i::nat. emeasure M (Q' i))" using a by simp
   let ?O = "\<lambda>n. \<Union>i\<le>n. Q' i"
-  have Union: "(SUP i. \<mu> (?O i)) = \<mu> (\<Union>i. ?O i)"
-  proof (rule continuity_from_below[of ?O])
-    show "range ?O \<subseteq> sets M" using Q' by (auto intro!: finite_UN)
+  have Union: "(SUP i. emeasure M (?O i)) = emeasure M (\<Union>i. ?O i)"
+  proof (rule SUP_emeasure_incseq[of ?O])
+    show "range ?O \<subseteq> sets M" using Q' by auto
     show "incseq ?O" by (fastforce intro!: incseq_SucI)
   qed
   have Q'_sets: "\<And>i. Q' i \<in> sets M" using Q' by auto
-  have O_sets: "\<And>i. ?O i \<in> sets M"
-     using Q' by (auto intro!: finite_UN Un)
+  have O_sets: "\<And>i. ?O i \<in> sets M" using Q' by auto
   then have O_in_G: "\<And>i. ?O i \<in> ?Q"
   proof (safe del: notI)
-    fix i have "Q' ` {..i} \<subseteq> sets M"
-      using Q' by (auto intro: finite_UN)
-    with v.measure_finitely_subadditive[of "{.. i}" Q']
-    have "\<nu> (?O i) \<le> (\<Sum>i\<le>i. \<nu> (Q' i))" by auto
+    fix i have "Q' ` {..i} \<subseteq> sets M" using Q' by auto
+    then have "N (?O i) \<le> (\<Sum>i\<le>i. N (Q' i))"
+      by (simp add: sets_eq emeasure_subadditive_finite)
     also have "\<dots> < \<infinity>" using Q' by (simp add: setsum_Pinfty)
-    finally show "\<nu> (?O i) \<noteq> \<infinity>" by simp
+    finally show "N (?O i) \<noteq> \<infinity>" by simp
   qed auto
   have O_mono: "\<And>n. ?O n \<subseteq> ?O (Suc n)" by fastforce
-  have a_eq: "?a = \<mu> (\<Union>i. ?O i)" unfolding Union[symmetric]
+  have a_eq: "?a = emeasure M (\<Union>i. ?O i)" unfolding Union[symmetric]
   proof (rule antisym)
-    show "?a \<le> (SUP i. \<mu> (?O i))" unfolding a_Lim
-      using Q' by (auto intro!: SUP_mono measure_mono finite_UN)
-    show "(SUP i. \<mu> (?O i)) \<le> ?a" unfolding SUP_def
+    show "?a \<le> (SUP i. emeasure M (?O i))" unfolding a_Lim
+      using Q' by (auto intro!: SUP_mono emeasure_mono)
+    show "(SUP i. emeasure M (?O i)) \<le> ?a" unfolding SUP_def
     proof (safe intro!: Sup_mono, unfold bex_simps)
       fix i
       have *: "(\<Union>Q' ` {..i}) = ?O i" by auto
-      then show "\<exists>x. (x \<in> sets M \<and> \<nu> x \<noteq> \<infinity>) \<and>
-        \<mu> (\<Union>Q' ` {..i}) \<le> \<mu> x"
+      then show "\<exists>x. (x \<in> sets M \<and> N x \<noteq> \<infinity>) \<and>
+        emeasure M (\<Union>Q' ` {..i}) \<le> emeasure M x"
         using O_in_G[of i] by (auto intro!: exI[of _ "?O i"])
     qed
   qed
@@ -687,51 +591,50 @@
     show "range Q \<subseteq> sets M"
       using Q_sets by auto
     { 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 = \<infinity>"
+      show "emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"
       proof (rule disjCI, simp)
-        assume *: "0 < \<mu> A \<longrightarrow> \<nu> A \<noteq> \<infinity>"
-        show "\<mu> A = 0 \<and> \<nu> A = 0"
+        assume *: "0 < emeasure M A \<longrightarrow> N A \<noteq> \<infinity>"
+        show "emeasure M A = 0 \<and> N A = 0"
         proof cases
-          assume "\<mu> A = 0" moreover with ac A have "\<nu> A = 0"
+          assume "emeasure M A = 0" moreover with ac A have "N A = 0"
             unfolding absolutely_continuous_def by auto
           ultimately show ?thesis by simp
         next
-          assume "\<mu> A \<noteq> 0" with * have "\<nu> A \<noteq> \<infinity>" using positive_measure[OF A(1)] 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])
+          assume "emeasure M A \<noteq> 0" with * have "N A \<noteq> \<infinity>" using emeasure_nonneg[of M A] by auto
+          with A have "emeasure M ?O_0 + emeasure M A = emeasure M (?O_0 \<union> A)"
+            using Q' by (auto intro!: plus_emeasure countable_UN)
+          also have "\<dots> = (SUP i. emeasure M (?O i \<union> A))"
+          proof (rule SUP_emeasure_incseq[of "\<lambda>i. ?O i \<union> A", symmetric, simplified])
             show "range (\<lambda>i. ?O i \<union> A) \<subseteq> sets M"
-              using `\<nu> A \<noteq> \<infinity>` O_sets A by auto
+              using `N A \<noteq> \<infinity>` O_sets A by auto
           qed (fastforce intro!: incseq_SucI)
           also have "\<dots> \<le> ?a"
           proof (safe intro!: SUP_least)
             fix i have "?O i \<union> A \<in> ?Q"
             proof (safe del: notI)
               show "?O i \<union> A \<in> sets M" using O_sets A by auto
-              from O_in_G[of i] have "\<nu> (?O i \<union> A) \<le> \<nu> (?O i) + \<nu> A"
-                using v.measure_subadditive[of "?O i" A] A O_sets by auto
-              with O_in_G[of i] show "\<nu> (?O i \<union> A) \<noteq> \<infinity>"
-                using `\<nu> A \<noteq> \<infinity>` by auto
+              from O_in_G[of i] have "N (?O i \<union> A) \<le> N (?O i) + N A"
+                using emeasure_subadditive[of "?O i" N A] A O_sets by (auto simp: sets_eq)
+              with O_in_G[of i] show "N (?O i \<union> A) \<noteq> \<infinity>"
+                using `N A \<noteq> \<infinity>` by auto
             qed
-            then show "\<mu> (?O i \<union> A) \<le> ?a" by (rule SUP_upper)
+            then show "emeasure M (?O i \<union> A) \<le> ?a" by (rule SUP_upper)
           qed
-          finally have "\<mu> A = 0"
-            unfolding a_eq using real_measure[OF `?O_0 \<in> sets M`] real_measure[OF A(1)] by auto
-          with `\<mu> A \<noteq> 0` show ?thesis by auto
+          finally have "emeasure M A = 0"
+            unfolding a_eq using measure_nonneg[of M A] by (simp add: emeasure_eq_measure)
+          with `emeasure M A \<noteq> 0` show ?thesis by auto
         qed
       qed }
-    { fix i show "\<nu> (Q i) \<noteq> \<infinity>"
+    { fix i show "N (Q i) \<noteq> \<infinity>"
       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`
-          using v.measure_mono[OF O_mono, of n] v.positive_measure[of "?O n"] v.positive_measure[of "?O (Suc n)"]
-          using v.measure_Diff[of "?O n" "?O (Suc n)", OF _ _ _ O_mono]
-          by (cases rule: ereal2_cases[of "\<nu> (\<Union> x\<le>Suc n. Q' x)" "\<nu> (\<Union> i\<le>n. Q' i)"]) auto
+        with `?O n \<in> ?Q` `?O (Suc n) \<in> ?Q`
+            emeasure_Diff[OF _ _ _ O_mono, of N n] emeasure_nonneg[of N "(\<Union> x\<le>n. Q' x)"]
+        show ?thesis
+          by (auto simp: sets_eq ereal_minus_eq_PInfty_iff Q_def)
       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)"
@@ -751,96 +654,93 @@
 qed
 
 lemma (in finite_measure) Radon_Nikodym_finite_measure_infinite:
-  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
-  assumes "absolutely_continuous \<nu>"
-  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
+  assumes "absolutely_continuous M N" and sets_eq: "sets N = sets M"
+  shows "\<exists>f\<in>borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"
 proof -
-  interpret v: measure_space ?N
-    where "space ?N = space M" and "sets ?N = sets M" and "measure ?N = \<nu>"
-    by fact auto
   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 = \<infinity>"
-    and Q_fin: "\<And>i. \<nu> (Q i) \<noteq> \<infinity>" by force
+    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> emeasure M A = 0 \<and> N A = 0 \<or> 0 < emeasure M A \<and> N A = \<infinity>"
+    and Q_fin: "\<And>i. N (Q i) \<noteq> \<infinity>" by force
   from Q have Q_sets: "\<And>i. Q i \<in> sets M" by auto
-  have "\<forall>i. \<exists>f. f\<in>borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M.
-    \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
-  proof
+  let ?N = "\<lambda>i. density N (indicator (Q i))" and ?M = "\<lambda>i. density M (indicator (Q i))"
+  have "\<forall>i. \<exists>f\<in>borel_measurable (?M i). (\<forall>x. 0 \<le> f x) \<and> density (?M i) f = ?N i"
+  proof (intro allI finite_measure.Radon_Nikodym_finite_measure)
     fix i
-    have indicator_eq: "\<And>f x A. (f x :: ereal) * indicator (Q i \<inter> A) x * indicator (Q i) x
-      = (f x * indicator (Q i) x) * indicator A x"
-      unfolding indicator_def by auto
-    have fm: "finite_measure (restricted_space (Q i))"
-      (is "finite_measure ?R") by (rule restricted_finite_measure[OF Q_sets[of i]])
-    then interpret R: finite_measure ?R .
-    have fmv: "finite_measure (restricted_space (Q i) \<lparr> measure := \<nu>\<rparr>)" (is "finite_measure ?Q")
-    proof
-      show "measure_space ?Q"
-        using v.restricted_measure_space Q_sets[of i] by auto
-      show "measure ?Q (space ?Q) \<noteq> \<infinity>" using Q_fin by simp
-    qed
-    have "R.absolutely_continuous \<nu>"
-      using `absolutely_continuous \<nu>` `Q i \<in> sets M`
-      by (auto simp: R.absolutely_continuous_def absolutely_continuous_def)
-    from R.Radon_Nikodym_finite_measure[OF fmv this]
-    obtain f where f: "(\<lambda>x. f x * indicator (Q i) x) \<in> borel_measurable M"
-      and f_int: "\<And>A. A\<in>sets M \<Longrightarrow> \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. (f x * indicator (Q i) x) * indicator A x \<partial>M)"
-      unfolding Bex_def borel_measurable_restricted[OF `Q i \<in> sets M`]
-        positive_integral_restricted[OF `Q i \<in> sets M`] by (auto simp: indicator_eq)
-    then show "\<exists>f. f\<in>borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M.
-      \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f x * indicator (Q i \<inter> A) x \<partial>M))"
-      by (auto intro!: exI[of _ "\<lambda>x. max 0 (f x * indicator (Q i) x)"] positive_integral_cong_pos
-        split: split_indicator split_if_asm simp: max_def)
+    from Q show "finite_measure (?M i)"
+      by (auto intro!: finite_measureI cong: positive_integral_cong
+               simp add: emeasure_density subset_eq sets_eq)
+    from Q have "emeasure (?N i) (space N) = emeasure N (Q i)"
+      by (simp add: sets_eq[symmetric] emeasure_density subset_eq cong: positive_integral_cong)
+    with Q_fin show "finite_measure (?N i)"
+      by (auto intro!: finite_measureI)
+    show "sets (?N i) = sets (?M i)" by (simp add: sets_eq)
+    show "absolutely_continuous (?M i) (?N i)"
+      using `absolutely_continuous M N` `Q i \<in> sets M`
+      by (auto simp: absolutely_continuous_def null_sets_density_iff sets_eq
+               intro!: absolutely_continuous_AE[OF sets_eq])
   qed
-  from choice[OF this] obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
-    and f: "\<And>A i. A \<in> sets M \<Longrightarrow>
-      \<nu> (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)"
+  from choice[OF this[unfolded Bex_def]]
+  obtain f where borel: "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x"
+    and f_density: "\<And>i. density (?M i) (f i) = ?N i"
     by auto
+  { fix A i assume A: "A \<in> sets M"
+    with Q borel have "(\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M) = emeasure (density (?M i) (f i)) A"
+      by (auto simp add: emeasure_density positive_integral_density subset_eq
+               intro!: positive_integral_cong split: split_indicator)
+    also have "\<dots> = emeasure N (Q i \<inter> A)"
+      using A Q by (simp add: f_density emeasure_restricted subset_eq sets_eq)
+    finally have "emeasure N (Q i \<inter> A) = (\<integral>\<^isup>+x. f i x * indicator (Q i \<inter> A) x \<partial>M)" .. }
+  note integral_eq = this
   let ?f = "\<lambda>x. (\<Sum>i. f i x * indicator (Q i) x) + \<infinity> * indicator Q0 x"
   show ?thesis
   proof (safe intro!: bexI[of _ ?f])
     show "?f \<in> borel_measurable M" using Q0 borel Q_sets
       by (auto intro!: measurable_If)
     show "\<And>x. 0 \<le> ?f x" using borel by (auto intro!: suminf_0_le simp: indicator_def)
-    fix A assume "A \<in> sets M"
-    have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
-    have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
-      "\<And>i. AE x. 0 \<le> f i x * indicator (Q i \<inter> A) x"
-      using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times)
-    have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"
-      using borel by (intro positive_integral_cong) (auto simp: indicator_def)
-    also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * \<mu> (Q0 \<inter> A)"
-      using borel Qi Q0(1) `A \<in> sets M`
-      by (subst positive_integral_add) (auto simp del: ereal_infty_mult
-          simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le)
-    also have "\<dots> = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)"
-      by (subst f[OF `A \<in> sets M`], subst positive_integral_suminf) auto
-    finally have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. \<nu> (Q i \<inter> A)) + \<infinity> * \<mu> (Q0 \<inter> A)" .
-    moreover have "(\<Sum>i. \<nu> (Q i \<inter> A)) = \<nu> ((\<Union>i. Q i) \<inter> A)"
-      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 "\<infinity> * \<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 "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 = (\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M)"
-      using v.measure_additive[simplified, of "(\<Union>i. Q i) \<inter> A" "Q0 \<inter> A"]
-      by simp
+    show "density M ?f = N"
+    proof (rule measure_eqI)
+      fix A assume "A \<in> sets (density M ?f)"
+      then have "A \<in> sets M" by simp
+      have Qi: "\<And>i. Q i \<in> sets M" using Q by auto
+      have [intro,simp]: "\<And>i. (\<lambda>x. f i x * indicator (Q i \<inter> A) x) \<in> borel_measurable M"
+        "\<And>i. AE x in M. 0 \<le> f i x * indicator (Q i \<inter> A) x"
+        using borel Qi Q0(1) `A \<in> sets M` by (auto intro!: borel_measurable_ereal_times)
+      have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) + \<infinity> * indicator (Q0 \<inter> A) x \<partial>M)"
+        using borel by (intro positive_integral_cong) (auto simp: indicator_def)
+      also have "\<dots> = (\<integral>\<^isup>+x. (\<Sum>i. f i x * indicator (Q i \<inter> A) x) \<partial>M) + \<infinity> * emeasure M (Q0 \<inter> A)"
+        using borel Qi Q0(1) `A \<in> sets M`
+        by (subst positive_integral_add) (auto simp del: ereal_infty_mult
+            simp add: positive_integral_cmult_indicator Int intro!: suminf_0_le)
+      also have "\<dots> = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)"
+        by (subst integral_eq[OF `A \<in> sets M`], subst positive_integral_suminf) auto
+      finally have "(\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M) = (\<Sum>i. N (Q i \<inter> A)) + \<infinity> * emeasure M (Q0 \<inter> A)" .
+      moreover have "(\<Sum>i. N (Q i \<inter> A)) = N ((\<Union>i. Q i) \<inter> A)"
+        using Q Q_sets `A \<in> sets M`
+        by (subst suminf_emeasure) (auto simp: disjoint_family_on_def sets_eq)
+      moreover have "\<infinity> * emeasure M (Q0 \<inter> A) = N (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 "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
+      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 have "N A = (\<integral>\<^isup>+x. ?f x * indicator A x \<partial>M)"
+        using plus_emeasure[of "(\<Union>i. Q i) \<inter> A" N "Q0 \<inter> A"] by (simp add: sets_eq)
+      with `A \<in> sets M` borel Q Q0(1) show "emeasure (density M ?f) A = N A"
+        by (subst emeasure_density)
+           (auto intro!: borel_measurable_ereal_add borel_measurable_psuminf measurable_If
+                 simp: subset_eq)
+    qed (simp add: sets_eq)
   qed
 qed
 
 lemma (in sigma_finite_measure) Radon_Nikodym:
-  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
-  assumes ac: "absolutely_continuous \<nu>"
-  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> (\<forall>A\<in>sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
+  assumes ac: "absolutely_continuous M N" assumes sets_eq: "sets N = sets M"
+  shows "\<exists>f \<in> borel_measurable M. (\<forall>x. 0 \<le> f x) \<and> density M f = N"
 proof -
   from Ex_finite_integrable_function
   obtain h where finite: "integral\<^isup>P M h \<noteq> \<infinity>" and
@@ -849,47 +749,38 @@
     pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x" and
     "\<And>x. x \<in> space M \<Longrightarrow> h x < \<infinity>" by auto
   let ?T = "\<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M)"
-  let ?MT = "M\<lparr> measure := ?T \<rparr>"
-  interpret T: finite_measure ?MT
-    where "space ?MT = space M" and "sets ?MT = sets M" and "measure ?MT = ?T"
-    using borel finite nn
-    by (auto intro!: measure_space_density finite_measureI cong: positive_integral_cong)
-  have "T.absolutely_continuous \<nu>"
-  proof (unfold T.absolutely_continuous_def, safe)
-    fix N assume "N \<in> sets M" "(\<integral>\<^isup>+x. h x * indicator N x \<partial>M) = 0"
-    with borel ac pos have "AE x. x \<notin> N"
-      by (subst (asm) positive_integral_0_iff_AE) (auto split: split_indicator simp: not_le)
-    then have "N \<in> null_sets" using `N \<in> sets M` sets_into_space
-      by (subst (asm) AE_iff_measurable[OF `N \<in> sets M`]) auto
-    then show "\<nu> N = 0" using ac by (auto simp: absolutely_continuous_def)
-  qed
-  from T.Radon_Nikodym_finite_measure_infinite[simplified, OF assms(1) this]
-  obtain f where f_borel: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and
-    fT: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+ x. f x * indicator A x \<partial>?MT)"
-    by (auto simp: measurable_def)
-  show ?thesis
-  proof (safe intro!: bexI[of _ "\<lambda>x. h x * f x"])
-    show "(\<lambda>x. h x * f x) \<in> borel_measurable M"
-      using borel f_borel by auto
-    show "\<And>x. 0 \<le> h x * f x" using nn f_borel by auto
-    fix A assume "A \<in> sets M"
-    then show "\<nu> A = (\<integral>\<^isup>+x. h x * f x * indicator A x \<partial>M)"
-      unfolding fT[OF `A \<in> sets M`] mult_assoc using nn borel f_borel
-      by (intro positive_integral_translated_density) auto
-  qed
+  let ?MT = "density M h"
+  from borel finite nn interpret T: finite_measure ?MT
+    by (auto intro!: finite_measureI cong: positive_integral_cong simp: emeasure_density)
+  have "absolutely_continuous ?MT N" "sets N = sets ?MT"
+  proof (unfold absolutely_continuous_def, safe)
+    fix A assume "A \<in> null_sets ?MT"
+    with borel have "A \<in> sets M" "AE x in M. x \<in> A \<longrightarrow> h x \<le> 0"
+      by (auto simp add: null_sets_density_iff)
+    with pos sets_into_space have "AE x in M. x \<notin> A"
+      by (elim eventually_elim1) (auto simp: not_le[symmetric])
+    then have "A \<in> null_sets M"
+      using `A \<in> sets M` by (simp add: AE_iff_null_sets)
+    with ac show "A \<in> null_sets N"
+      by (auto simp: absolutely_continuous_def)
+  qed (auto simp add: sets_eq)
+  from T.Radon_Nikodym_finite_measure_infinite[OF this]
+  obtain f where f_borel: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "density ?MT f = N" by auto
+  with nn borel show ?thesis
+    by (auto intro!: bexI[of _ "\<lambda>x. h x * f x"] simp: density_density_eq)
 qed
 
 section "Uniqueness of densities"
 
-lemma (in measure_space) finite_density_unique:
+lemma finite_density_unique:
   assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  assumes pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> g x"
+  assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
   and fin: "integral\<^isup>P M f \<noteq> \<infinity>"
   shows "(\<forall>A\<in>sets M. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. g x * indicator A x \<partial>M))
-    \<longleftrightarrow> (AE x. f x = g x)"
+    \<longleftrightarrow> (AE x in M. 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"
+  fix A assume eq: "AE x in M. f x = g x"
   then show "?P f A = ?P g A"
     by (auto intro: positive_integral_cong_AE)
 next
@@ -897,7 +788,7 @@
   from this[THEN bspec, OF top] fin
   have g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" by (simp cong: positive_integral_cong)
   { fix f g assume borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-      and pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> g x"
+      and pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> g x"
       and g_fin: "integral\<^isup>P M g \<noteq> \<infinity>" and eq: "\<forall>A\<in>sets M. ?P f A = ?P g A"
     let ?N = "{x\<in>space M. g x < f x}"
     have N: "?N \<in> sets M" using borel by simp
@@ -910,167 +801,163 @@
     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
-      show "AE x. g x * indicator ?N x \<le> f x * indicator ?N x"
-           "AE x. 0 \<le> g x * indicator ?N x"
+      show "AE x in M. g x * indicator ?N x \<le> f x * indicator ?N x"
+           "AE x in M. 0 \<le> g x * indicator ?N x"
         using pos by (auto split: split_indicator)
     qed fact
     also have "\<dots> = 0"
-      unfolding eq[THEN bspec, OF N] using positive_integral_positive Pg_fin by auto
-    finally have "AE x. f x \<le> g x"
+      unfolding eq[THEN bspec, OF N] using positive_integral_positive[of M] Pg_fin by auto
+    finally have "AE x in M. f x \<le> g x"
       using pos borel positive_integral_PInf_AE[OF borel(2) g_fin]
       by (subst (asm) positive_integral_0_iff_AE)
          (auto split: split_indicator simp: not_less ereal_minus_le_iff) }
   from this[OF borel pos g_fin eq] this[OF borel(2,1) pos(2,1) fin] eq
-  show "AE x. f x = g x" by auto
+  show "AE x in M. f x = g x" by auto
 qed
 
 lemma (in finite_measure) density_unique_finite_measure:
   assumes borel: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"
-  assumes pos: "AE x. 0 \<le> f x" "AE x. 0 \<le> f' x"
+  assumes pos: "AE x in M. 0 \<le> f x" "AE x in M. 0 \<le> f' x"
   assumes f: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)"
     (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
-  shows "AE x. f x = f' x"
+  shows "AE x in M. f x = f' x"
 proof -
-  let ?\<nu> = "\<lambda>A. ?P f A" and ?\<nu>' = "\<lambda>A. ?P f' A"
+  let ?D = "\<lambda>f. density M f"
+  let ?N = "\<lambda>A. ?P f A" and ?N' = "\<lambda>A. ?P f' A"
   let ?f = "\<lambda>A x. f x * indicator A x" and ?f' = "\<lambda>A x. f' x * indicator A x"
-  interpret M: measure_space "M\<lparr> measure := ?\<nu>\<rparr>"
-    using borel(1) pos(1) by (rule measure_space_density) simp
-  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\<lparr> measure := ?\<nu>\<rparr>)` ac]
+
+  have ac: "absolutely_continuous M (density M f)" "sets (density M f) = sets M"
+    using borel by (auto intro!: absolutely_continuousI_density) 
+  from split_space_into_finite_sets_and_rest[OF this]
   obtain 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 = \<infinity>"
-    and Q_fin: "\<And>i. ?\<nu> (Q i) \<noteq> \<infinity>" by force
+    and in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> emeasure M A = 0 \<and> ?D f A = 0 \<or> 0 < emeasure M A \<and> ?D f A = \<infinity>"
+    and Q_fin: "\<And>i. ?D f (Q i) \<noteq> \<infinity>" by force
+  with borel pos have in_Q0: "\<And>A. A \<in> sets M \<Longrightarrow> A \<subseteq> Q0 \<Longrightarrow> emeasure M A = 0 \<and> ?N A = 0 \<or> 0 < emeasure M A \<and> ?N A = \<infinity>"
+    and Q_fin: "\<And>i. ?N (Q i) \<noteq> \<infinity>" by (auto simp: emeasure_density subset_eq)
+
   from Q have Q_sets: "\<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
+  let ?D = "{x\<in>space M. f x \<noteq> f' x}"
+  have "?D \<in> sets M" using borel by auto
   have *: "\<And>i x A. \<And>y::ereal. y * indicator (Q i) x * indicator A x = y * indicator (Q i \<inter> A) x"
     unfolding indicator_def by auto
-  have "\<forall>i. AE x. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos
+  have "\<forall>i. AE x in M. ?f (Q i) x = ?f' (Q i) x" using borel Q_fin Q pos
     by (intro finite_density_unique[THEN iffD1] allI)
        (auto intro!: borel_measurable_ereal_times f Int simp: *)
-  moreover have "AE x. ?f Q0 x = ?f' Q0 x"
+  moreover have "AE x in M. ?f Q0 x = ?f' Q0 x"
   proof (rule AE_I')
     { fix f :: "'a \<Rightarrow> ereal" assume borel: "f \<in> borel_measurable M"
-        and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?\<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
+        and eq: "\<And>A. A \<in> sets M \<Longrightarrow> ?N A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
       let ?A = "\<lambda>i. Q0 \<inter> {x \<in> space M. f x < (i::nat)}"
-      have "(\<Union>i. ?A i) \<in> null_sets"
+      have "(\<Union>i. ?A i) \<in> null_sets M"
       proof (rule null_sets_UN)
         fix i ::nat have "?A i \<in> sets M"
           using borel Q0(1) by auto
-        have "?\<nu> (?A i) \<le> (\<integral>\<^isup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"
+        have "?N (?A i) \<le> (\<integral>\<^isup>+x. (i::ereal) * indicator (?A i) x \<partial>M)"
           unfolding eq[OF `?A i \<in> sets M`]
           by (auto intro!: positive_integral_mono simp: indicator_def)
-        also have "\<dots> = i * \<mu> (?A i)"
+        also have "\<dots> = i * emeasure M (?A i)"
           using `?A i \<in> sets M` by (auto intro!: positive_integral_cmult_indicator)
-        also have "\<dots> < \<infinity>"
-          using `?A i \<in> sets M`[THEN finite_measure] by auto
-        finally have "?\<nu> (?A i) \<noteq> \<infinity>" 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
+        also have "\<dots> < \<infinity>" using emeasure_real[of "?A i"] by simp
+        finally have "?N (?A i) \<noteq> \<infinity>" by simp
+        then show "?A i \<in> null_sets M" 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 \<noteq> \<infinity>}"
         by (auto simp: less_PInf_Ex_of_nat real_eq_of_nat)
-      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets" by simp }
+      finally have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" by simp }
     from this[OF borel(1) refl] this[OF borel(2) f]
-    have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets" by simp_all
-    then show "(Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>}) \<in> null_sets" by (rule nullsets.Un)
+    have "Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>} \<in> null_sets M" "Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>} \<in> null_sets M" by simp_all
+    then show "(Q0 \<inter> {x\<in>space M. f x \<noteq> \<infinity>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>}) \<in> null_sets M" 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> \<infinity>}) \<union> (Q0 \<inter> {x\<in>space M. f' x \<noteq> \<infinity>})" by (auto simp: indicator_def)
   qed
-  moreover have "\<And>x. (?f Q0 x = ?f' Q0 x) \<longrightarrow> (\<forall>i. ?f (Q i) x = ?f' (Q i) x) \<longrightarrow>
+  moreover have "AE x in M. (?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)
-  ultimately have "AE x. ?f (space M) x = ?f' (space M) x"
-    by (auto simp: AE_all_countable[symmetric])
-  then show "AE x. f x = f' x" by auto
+  ultimately have "AE x in M. ?f (space M) x = ?f' (space M) x"
+    unfolding AE_all_countable[symmetric]
+    by eventually_elim (auto intro!: AE_I2 split: split_if_asm simp: indicator_def)
+  then show "AE x in M. f x = f' x" by auto
 qed
 
 lemma (in sigma_finite_measure) density_unique:
-  assumes f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
-  assumes f': "f' \<in> borel_measurable M" "AE x. 0 \<le> f' x"
-  assumes eq: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)"
-    (is "\<And>A. A \<in> sets M \<Longrightarrow> ?P f A = ?P f' A")
-  shows "AE x. f x = f' x"
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  assumes f': "f' \<in> borel_measurable M" "AE x in M. 0 \<le> f' x"
+  assumes density_eq: "density M f = density M f'"
+  shows "AE x in M. f x = f' x"
 proof -
   obtain h where h_borel: "h \<in> borel_measurable M"
     and fin: "integral\<^isup>P M h \<noteq> \<infinity>" and pos: "\<And>x. x \<in> space M \<Longrightarrow> 0 < h x \<and> h x < \<infinity>" "\<And>x. 0 \<le> h x"
     using Ex_finite_integrable_function by auto
-  then have h_nn: "AE x. 0 \<le> h x" by auto
-  let ?H = "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>"
-  have H: "measure_space ?H"
-    using h_borel h_nn by (rule measure_space_density) simp
-  then interpret h: measure_space ?H .
-  interpret h: finite_measure "M\<lparr> measure := \<lambda>A. (\<integral>\<^isup>+x. h x * indicator A x \<partial>M) \<rparr>"
-    by (intro H finite_measureI) (simp cong: positive_integral_cong add: fin)
-  let ?fM = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)\<rparr>"
-  interpret f: measure_space ?fM
-    using f by (rule measure_space_density) simp
-  let ?f'M = "M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. f' x * indicator A x \<partial>M)\<rparr>"
-  interpret f': measure_space ?f'M
-    using f' by (rule measure_space_density) simp
+  then have h_nn: "AE x in M. 0 \<le> h x" by auto
+  let ?H = "density M h"
+  interpret h: finite_measure ?H
+    using fin h_borel pos
+    by (intro finite_measureI) (simp cong: positive_integral_cong emeasure_density add: fin)
+  let ?fM = "density M f"
+  let ?f'M = "density M f'"
   { fix A assume "A \<in> sets M"
     then have "{x \<in> space M. h x * indicator A x \<noteq> 0} = A"
       using pos(1) sets_into_space by (force simp: indicator_def)
-    then have "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets"
+    then have "(\<integral>\<^isup>+x. h x * indicator A x \<partial>M) = 0 \<longleftrightarrow> A \<in> null_sets M"
       using h_borel `A \<in> sets M` h_nn by (subst positive_integral_0_iff) auto }
   note h_null_sets = this
   { fix A assume "A \<in> sets M"
     have "(\<integral>\<^isup>+x. f x * (h x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?fM)"
       using `A \<in> sets M` h_borel h_nn f f'
-      by (intro positive_integral_translated_density[symmetric]) auto
+      by (intro positive_integral_density[symmetric]) auto
     also have "\<dots> = (\<integral>\<^isup>+x. h x * indicator A x \<partial>?f'M)"
-      by (rule f'.positive_integral_cong_measure) (simp_all add: eq)
+      by (simp_all add: density_eq)
     also have "\<dots> = (\<integral>\<^isup>+x. f' x * (h x * indicator A x) \<partial>M)"
       using `A \<in> sets M` h_borel h_nn f f'
-      by (intro positive_integral_translated_density) auto
+      by (intro positive_integral_density) auto
     finally have "(\<integral>\<^isup>+x. h x * (f x * indicator A x) \<partial>M) = (\<integral>\<^isup>+x. h x * (f' x * indicator A x) \<partial>M)"
       by (simp add: ac_simps)
     then have "(\<integral>\<^isup>+x. (f x * indicator A x) \<partial>?H) = (\<integral>\<^isup>+x. (f' x * indicator A x) \<partial>?H)"
       using `A \<in> sets M` h_borel h_nn f f'
-      by (subst (asm) (1 2) positive_integral_translated_density[symmetric]) auto }
+      by (subst (asm) (1 2) positive_integral_density[symmetric]) auto }
   then have "AE x in ?H. f x = f' x" using h_borel h_nn f f'
-    by (intro h.density_unique_finite_measure absolutely_continuous_AE[OF H] density_is_absolutely_continuous)
-       simp_all
-  then show "AE x. f x = f' x"
-    unfolding h.almost_everywhere_def almost_everywhere_def
-    by (auto simp add: h_null_sets)
+    by (intro h.density_unique_finite_measure absolutely_continuous_AE[of M])
+       (auto simp add: AE_density)
+  then show "AE x in M. f x = f' x"
+    unfolding eventually_ae_filter using h_borel pos
+    by (auto simp add: h_null_sets null_sets_density_iff not_less[symmetric]
+                          AE_iff_null_sets[symmetric])
+       blast
 qed
 
+lemma (in sigma_finite_measure) density_unique_iff:
+  assumes f: "f \<in> borel_measurable M" and "AE x in M. 0 \<le> f x"
+  assumes f': "f' \<in> borel_measurable M" and "AE x in M. 0 \<le> f' x"
+  shows "density M f = density M f' \<longleftrightarrow> (AE x in M. f x = f' x)"
+  using density_unique[OF assms] density_cong[OF f f'] by auto
+
 lemma (in sigma_finite_measure) sigma_finite_iff_density_finite:
-  assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" (is "measure_space ?N")
-    and f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
-    and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
-  shows "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>) \<longleftrightarrow> (AE x. f x \<noteq> \<infinity>)"
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  shows "sigma_finite_measure (density M f) \<longleftrightarrow> (AE x in M. f x \<noteq> \<infinity>)"
+    (is "sigma_finite_measure ?N \<longleftrightarrow> _")
 proof
   assume "sigma_finite_measure ?N"
-  then interpret \<nu>: sigma_finite_measure ?N
-    where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
-    and "sets ?N = sets M" and "space ?N = space M" by (auto simp: measurable_def)
-  from \<nu>.Ex_finite_integrable_function obtain h where
+  then interpret N: sigma_finite_measure ?N .
+  from N.Ex_finite_integrable_function obtain h where
     h: "h \<in> borel_measurable M" "integral\<^isup>P ?N h \<noteq> \<infinity>" and
     h_nn: "\<And>x. 0 \<le> h x" and
     fin: "\<forall>x\<in>space M. 0 < h x \<and> h x < \<infinity>" by auto
-  have "AE x. f x * h x \<noteq> \<infinity>"
+  have "AE x in M. f x * h x \<noteq> \<infinity>"
   proof (rule AE_I')
     have "integral\<^isup>P ?N h = (\<integral>\<^isup>+x. f x * h x \<partial>M)" using f h h_nn
-      by (subst \<nu>.positive_integral_cong_measure[symmetric,
-          of "M\<lparr> measure := \<lambda> A. \<integral>\<^isup>+x. f x * indicator A x \<partial>M \<rparr>"])
-         (auto intro!: positive_integral_translated_density simp: eq)
+      by (auto intro!: positive_integral_density)
     then have "(\<integral>\<^isup>+x. f x * h x \<partial>M) \<noteq> \<infinity>"
       using h(2) by simp
-    then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets"
+    then show "(\<lambda>x. f x * h x) -` {\<infinity>} \<inter> space M \<in> null_sets M"
       using f h(1) by (auto intro!: positive_integral_PInf borel_measurable_vimage)
   qed auto
-  then show "AE x. f x \<noteq> \<infinity>"
+  then show "AE x in M. f x \<noteq> \<infinity>"
     using fin by (auto elim!: AE_Ball_mp)
 next
-  assume AE: "AE x. f x \<noteq> \<infinity>"
+  assume AE: "AE x in M. f x \<noteq> \<infinity>"
   from sigma_finite guess Q .. note Q = this
-  interpret \<nu>: measure_space ?N
-    where "borel_measurable ?N = borel_measurable M" and "measure ?N = \<nu>"
-    and "sets ?N = sets M" and "space ?N = space M" using \<nu> by (auto simp: measurable_def)
   def A \<equiv> "\<lambda>i. f -` (case i of 0 \<Rightarrow> {\<infinity>} | Suc n \<Rightarrow> {.. ereal(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
@@ -1113,7 +1000,7 @@
     have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<noteq> \<infinity>"
     proof (cases i)
       case 0
-      have "AE x. f x * indicator (A i \<inter> Q j) x = 0"
+      have "AE x in M. f x * indicator (A i \<inter> Q j) x = 0"
         using AE by (auto simp: A_def `i = 0`)
       from positive_integral_cong_AE[OF this] show ?thesis by simp
     next
@@ -1121,271 +1008,238 @@
       then have "(\<integral>\<^isup>+x. f x * indicator (A i \<inter> Q j) x \<partial>M) \<le>
         (\<integral>\<^isup>+x. (Suc n :: ereal) * indicator (Q j) x \<partial>M)"
         by (auto intro!: positive_integral_mono simp: indicator_def A_def real_eq_of_nat)
-      also have "\<dots> = Suc n * \<mu> (Q j)"
+      also have "\<dots> = Suc n * emeasure M (Q j)"
         using Q by (auto intro!: positive_integral_cmult_indicator)
       also have "\<dots> < \<infinity>"
         using Q by (auto simp: real_eq_of_nat[symmetric])
       finally show ?thesis by simp
     qed
-    then show "measure ?N (?A n) \<noteq> \<infinity>"
-      using A_in_sets Q eq by auto
+    then show "emeasure ?N (?A n) \<noteq> \<infinity>"
+      using A_in_sets Q f by (auto simp: emeasure_density)
   qed
 qed
 
 section "Radon-Nikodym derivative"
 
 definition
-  "RN_deriv M \<nu> \<equiv> SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and>
-    (\<forall>A \<in> sets M. \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M))"
+  "RN_deriv M N \<equiv> SOME f. f \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> f x) \<and> density M f = N"
 
-lemma (in sigma_finite_measure) RN_deriv_cong:
-  assumes cong: "\<And>A. A \<in> sets M \<Longrightarrow> measure M' A = \<mu> A" "sets M' = sets M" "space M' = space M"
-    and \<nu>: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu>' A = \<nu> A"
-  shows "RN_deriv M' \<nu>' x = RN_deriv M \<nu> x"
+lemma
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  shows borel_measurable_RN_deriv_density: "RN_deriv M (density M f) \<in> borel_measurable M" (is ?borel)
+    and density_RN_deriv_density: "density M (RN_deriv M (density M f)) = density M f" (is ?density)
+    and RN_deriv_density_nonneg: "0 \<le> RN_deriv M (density M f) x" (is ?pos)
 proof -
-  interpret \<mu>': sigma_finite_measure M'
-    using cong by (rule sigma_finite_measure_cong)
-  show ?thesis
-    unfolding RN_deriv_def
-    by (simp add: cong \<nu> positive_integral_cong_measure[OF cong] measurable_def)
+  let ?f = "\<lambda>x. max 0 (f x)"
+  let ?P = "\<lambda>g. g \<in> borel_measurable M \<and> (\<forall>x. 0 \<le> g x) \<and> density M g = density M f"
+  from f have "?P ?f" using f by (auto intro!: density_cong simp: split: split_max)
+  then have "?P (RN_deriv M (density M f))"
+    unfolding RN_deriv_def by (rule someI[where P="?P"])
+  then show ?borel ?density ?pos by auto
 qed
 
 lemma (in sigma_finite_measure) RN_deriv:
-  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
-  assumes "absolutely_continuous \<nu>"
-  shows "RN_deriv M \<nu> \<in> borel_measurable M" (is ?borel)
-  and "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)"
-    (is "\<And>A. _ \<Longrightarrow> ?int A")
-  and "0 \<le> RN_deriv M \<nu> x"
+  assumes "absolutely_continuous M N" "sets N = sets M"
+  shows borel_measurable_RN_deriv: "RN_deriv M N \<in> borel_measurable M" (is ?borel)
+    and density_RN_deriv: "density M (RN_deriv M N) = N" (is ?density)
+    and RN_deriv_nonneg: "0 \<le> RN_deriv M N x" (is ?pos)
 proof -
   note Ex = Radon_Nikodym[OF assms, unfolded Bex_def]
-  then show ?borel unfolding RN_deriv_def by (rule someI2_ex) auto
-  from Ex show "0 \<le> RN_deriv M \<nu> x" unfolding RN_deriv_def
-    by (rule someI2_ex) simp
-  fix A assume "A \<in> sets M"
-  from Ex show "?int A" unfolding RN_deriv_def
-    by (rule someI2_ex) (simp add: `A \<in> sets M`)
+  from Ex show ?borel unfolding RN_deriv_def by (rule someI2_ex) simp
+  from Ex show ?density unfolding RN_deriv_def by (rule someI2_ex) simp
+  from Ex show ?pos unfolding RN_deriv_def by (rule someI2_ex) simp
 qed
 
 lemma (in sigma_finite_measure) RN_deriv_positive_integral:
-  assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
+  assumes N: "absolutely_continuous M N" "sets N = sets M"
     and f: "f \<in> borel_measurable M"
-  shows "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
+  shows "integral\<^isup>P N f = (\<integral>\<^isup>+x. RN_deriv M N x * f x \<partial>M)"
 proof -
-  interpret \<nu>: measure_space "M\<lparr>measure := \<nu>\<rparr>" by fact
-  note RN = RN_deriv[OF \<nu>]
-  have "integral\<^isup>P (M\<lparr>measure := \<nu>\<rparr>) f = (\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<nu>\<rparr>)"
-    unfolding positive_integral_max_0 ..
-  also have "(\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<nu>\<rparr>) =
-    (\<integral>\<^isup>+x. max 0 (f x) \<partial>M\<lparr>measure := \<lambda>A. (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)\<rparr>)"
-    by (intro \<nu>.positive_integral_cong_measure[symmetric]) (simp_all add: RN(2))
-  also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * max 0 (f x) \<partial>M)"
-    by (intro positive_integral_translated_density) (auto simp add: RN f)
-  also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * f x \<partial>M)"
-    using RN_deriv(3)[OF \<nu>]
-    by (auto intro!: positive_integral_cong_pos split: split_if_asm
-             simp: max_def ereal_mult_le_0_iff)
-  finally show ?thesis .
-qed
-
-lemma (in sigma_finite_measure) RN_deriv_unique:
-  assumes \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
-  and f: "f \<in> borel_measurable M" "AE x. 0 \<le> f x"
-  and eq: "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = (\<integral>\<^isup>+x. f x * indicator A x \<partial>M)"
-  shows "AE x. f x = RN_deriv M \<nu> x"
-proof (rule density_unique[OF f RN_deriv(1)[OF \<nu>]])
-  show "AE x. 0 \<le> RN_deriv M \<nu> x" using RN_deriv[OF \<nu>] by auto
-  fix A assume A: "A \<in> sets M"
-  show "(\<integral>\<^isup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * indicator A x \<partial>M)"
-    unfolding eq[OF A, symmetric] RN_deriv(2)[OF \<nu> A, symmetric] ..
+  have "integral\<^isup>P N f = integral\<^isup>P (density M (RN_deriv M N)) f"
+    using N by (simp add: density_RN_deriv)
+  also have "\<dots> = (\<integral>\<^isup>+x. RN_deriv M N x * f x \<partial>M)"
+    using RN_deriv(1,3)[OF N] f by (simp add: positive_integral_density)
+  finally show ?thesis by simp
 qed
 
-lemma (in sigma_finite_measure) RN_deriv_vimage:
-  assumes T: "T \<in> measure_preserving M M'"
-    and T': "T' \<in> measure_preserving (M'\<lparr> measure := \<nu>' \<rparr>) (M\<lparr> measure := \<nu> \<rparr>)"
-    and inv: "\<And>x. x \<in> space M \<Longrightarrow> T' (T x) = x"
-  and \<nu>': "measure_space (M'\<lparr>measure := \<nu>'\<rparr>)" "measure_space.absolutely_continuous M' \<nu>'"
-  shows "AE x. RN_deriv M' \<nu>' (T x) = RN_deriv M \<nu> x"
+lemma null_setsD_AE: "N \<in> null_sets M \<Longrightarrow> AE x in M. x \<notin> N"
+  using AE_iff_null_sets[of N M] by auto
+
+lemma (in sigma_finite_measure) RN_deriv_unique:
+  assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x"
+  and eq: "density M f = N"
+  shows "AE x in M. f x = RN_deriv M N x"
+proof (rule density_unique)
+  have ac: "absolutely_continuous M N"
+    using f(1) unfolding eq[symmetric] by (rule absolutely_continuousI_density)
+  have eq2: "sets N = sets M"
+    unfolding eq[symmetric] by simp
+  show "RN_deriv M N \<in> borel_measurable M" "AE x in M. 0 \<le> RN_deriv M N x"
+    "density M f = density M (RN_deriv M N)"
+    using RN_deriv[OF ac eq2] eq by auto
+qed fact+
+
+lemma (in sigma_finite_measure) RN_deriv_distr:
+  fixes T :: "'a \<Rightarrow> 'b"
+  assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M"
+    and inv: "\<forall>x\<in>space M. T' (T x) = x"
+  and ac: "absolutely_continuous (distr M M' T) (distr N M' T)"
+  and N: "sets N = sets M"
+  shows "AE x in M. RN_deriv (distr M M' T) (distr N M' T) (T x) = RN_deriv M N x"
 proof (rule RN_deriv_unique)
-  interpret \<nu>': measure_space "M'\<lparr>measure := \<nu>'\<rparr>" by fact
-  show "measure_space (M\<lparr> measure := \<nu>\<rparr>)"
-    by (rule \<nu>'.measure_space_vimage[OF _ T'], simp) default
-  interpret M': measure_space M'
-  proof (rule measure_space_vimage)
-    have "sigma_algebra (M'\<lparr> measure := \<nu>'\<rparr>)" by default
-    then show "sigma_algebra M'" by simp
-  qed fact
-  show "absolutely_continuous \<nu>" unfolding absolutely_continuous_def
-  proof safe
-    fix N assume N: "N \<in> sets M" and N_0: "\<mu> N = 0"
-    then have N': "T' -` N \<inter> space M' \<in> sets M'"
-      using T' by (auto simp: measurable_def measure_preserving_def)
-    have "T -` (T' -` N \<inter> space M') \<inter> space M = N"
-      using inv T N sets_into_space[OF N] by (auto simp: measurable_def measure_preserving_def)
-    then have "measure M' (T' -` N \<inter> space M') = 0"
-      using measure_preservingD[OF T N'] N_0 by auto
-    with \<nu>'(2) N' show "\<nu> N = 0" using measure_preservingD[OF T', of N] N
-      unfolding M'.absolutely_continuous_def measurable_def by auto
-  qed
-  interpret M': sigma_finite_measure M'
+  have [simp]: "sets N = sets M" by fact
+  note sets_eq_imp_space_eq[OF N, simp]
+  have measurable_N[simp]: "\<And>M'. measurable N M' = measurable M M'" by (auto simp: measurable_def)
+  { fix A assume "A \<in> sets M"
+    with inv T T' sets_into_space[OF this]
+    have "T -` T' -` A \<inter> T -` space M' \<inter> space M = A"
+      by (auto simp: measurable_def) }
+  note eq = this[simp]
+  { fix A assume "A \<in> sets M"
+    with inv T T' sets_into_space[OF this]
+    have "(T' \<circ> T) -` A \<inter> space M = A"
+      by (auto simp: measurable_def) }
+  note eq2 = this[simp]
+  let ?M' = "distr M M' T" and ?N' = "distr N M' T"
+  interpret M': sigma_finite_measure ?M'
   proof
     from sigma_finite guess F .. note F = this
-    show "\<exists>A::nat \<Rightarrow> 'c set. range A \<subseteq> sets M' \<and> (\<Union>i. A i) = space M' \<and> (\<forall>i. M'.\<mu> (A i) \<noteq> \<infinity>)"
+    show "\<exists>A::nat \<Rightarrow> 'b set. range A \<subseteq> sets ?M' \<and> (\<Union>i. A i) = space ?M' \<and> (\<forall>i. emeasure ?M' (A i) \<noteq> \<infinity>)"
     proof (intro exI conjI allI)
-      show *: "range (\<lambda>i. T' -` F i \<inter> space M') \<subseteq> sets M'"
-        using F T' by (auto simp: measurable_def measure_preserving_def)
-      show "(\<Union>i. T' -` F i \<inter> space M') = space M'"
-        using F T' by (force simp: measurable_def measure_preserving_def)
+      show *: "range (\<lambda>i. T' -` F i \<inter> space ?M') \<subseteq> sets ?M'"
+        using F T' by (auto simp: measurable_def)
+      show "(\<Union>i. T' -` F i \<inter> space ?M') = space ?M'"
+        using F T' by (force simp: measurable_def)
       fix i
       have "T' -` F i \<inter> space M' \<in> sets M'" using * by auto
-      note measure_preservingD[OF T this, symmetric]
       moreover
       have Fi: "F i \<in> sets M" using F by auto
-      then have "T -` (T' -` F i \<inter> space M') \<inter> space M = F i"
-        using T inv sets_into_space[OF Fi]
-        by (auto simp: measurable_def measure_preserving_def)
-      ultimately show "measure M' (T' -` F i \<inter> space M') \<noteq> \<infinity>"
-        using F by simp
+      ultimately show "emeasure ?M' (T' -` F i \<inter> space ?M') \<noteq> \<infinity>"
+        using F T T' by (simp add: emeasure_distr)
     qed
   qed
-  have "(RN_deriv M' \<nu>') \<circ> T \<in> borel_measurable M"
-    by (intro measurable_comp[where b=M'] M'.RN_deriv(1) measure_preservingD2[OF T]) fact+
-  then show "(\<lambda>x. RN_deriv M' \<nu>' (T x)) \<in> borel_measurable M"
+  have "(RN_deriv ?M' ?N') \<circ> T \<in> borel_measurable M"
+    using T ac by (intro measurable_comp[where b="?M'"] M'.borel_measurable_RN_deriv) simp_all
+  then show "(\<lambda>x. RN_deriv ?M' ?N' (T x)) \<in> borel_measurable M"
     by (simp add: comp_def)
-  show "AE x. 0 \<le> RN_deriv M' \<nu>' (T x)" using M'.RN_deriv(3)[OF \<nu>'] by auto
-  fix A let ?A = "T' -` A \<inter> space M'"
-  assume A: "A \<in> sets M"
-  then have A': "?A \<in> sets M'" using T' unfolding measurable_def measure_preserving_def
-    by auto
-  from A have "\<nu> A = \<nu>' ?A" using T'[THEN measure_preservingD] by simp
-  also have "\<dots> = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' x * indicator ?A x \<partial>M'"
-    using A' by (rule M'.RN_deriv(2)[OF \<nu>'])
-  also have "\<dots> = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' (T x) * indicator ?A (T x) \<partial>M"
-  proof (rule positive_integral_vimage)
-    show "sigma_algebra M'" by default
-    show "(\<lambda>x. RN_deriv M' \<nu>' x * indicator (T' -` A \<inter> space M') x) \<in> borel_measurable M'"
-      by (auto intro!: A' M'.RN_deriv(1)[OF \<nu>'])
-  qed fact
-  also have "\<dots> = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' (T x) * indicator A x \<partial>M"
-    using T inv by (auto intro!: positive_integral_cong simp: measure_preserving_def measurable_def indicator_def)
-  finally show "\<nu> A = \<integral>\<^isup>+ x. RN_deriv M' \<nu>' (T x) * indicator A x \<partial>M" .
+  show "AE x in M. 0 \<le> RN_deriv ?M' ?N' (T x)" using M'.RN_deriv_nonneg[OF ac] by auto
+
+  have "N = distr N M (T' \<circ> T)"
+    by (subst measure_of_of_measure[of N, symmetric])
+       (auto simp add: distr_def sigma_sets_eq intro!: measure_of_eq space_closed)
+  also have "\<dots> = distr (distr N M' T) M T'"
+    using T T' by (simp add: distr_distr)
+  also have "\<dots> = distr (density (distr M M' T) (RN_deriv (distr M M' T) (distr N M' T))) M T'"
+    using ac by (simp add: M'.density_RN_deriv)
+  also have "\<dots> = density M (RN_deriv (distr M M' T) (distr N M' T) \<circ> T)"
+    using M'.borel_measurable_RN_deriv[OF ac] by (simp add: distr_density_distr[OF T T', OF inv])
+  finally show "density M (\<lambda>x. RN_deriv (distr M M' T) (distr N M' T) (T x)) = N"
+    by (simp add: comp_def)
 qed
 
 lemma (in sigma_finite_measure) RN_deriv_finite:
-  assumes sfm: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" and ac: "absolutely_continuous \<nu>"
-  shows "AE x. RN_deriv M \<nu> x \<noteq> \<infinity>"
+  assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"
+  shows "AE x in M. RN_deriv M N x \<noteq> \<infinity>"
 proof -
-  interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
-  have \<nu>: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
-  from sfm show ?thesis
-    using sigma_finite_iff_density_finite[OF \<nu> RN_deriv(1)[OF \<nu> ac]] RN_deriv(2,3)[OF \<nu> ac] by simp
+  interpret N: sigma_finite_measure N by fact
+  from N show ?thesis
+    using sigma_finite_iff_density_finite[OF RN_deriv(1)[OF ac]] RN_deriv(2,3)[OF ac] by simp
 qed
 
 lemma (in sigma_finite_measure)
-  assumes \<nu>: "sigma_finite_measure (M\<lparr>measure := \<nu>\<rparr>)" "absolutely_continuous \<nu>"
+  assumes N: "sigma_finite_measure N" and ac: "absolutely_continuous M N" "sets N = sets M"
     and f: "f \<in> borel_measurable M"
-  shows RN_deriv_integrable: "integrable (M\<lparr>measure := \<nu>\<rparr>) f \<longleftrightarrow>
-      integrable M (\<lambda>x. real (RN_deriv M \<nu> x) * f x)" (is ?integrable)
-    and RN_deriv_integral: "integral\<^isup>L (M\<lparr>measure := \<nu>\<rparr>) f =
-      (\<integral>x. real (RN_deriv M \<nu> x) * f x \<partial>M)" (is ?integral)
+  shows RN_deriv_integrable: "integrable N f \<longleftrightarrow>
+      integrable M (\<lambda>x. real (RN_deriv M N x) * f x)" (is ?integrable)
+    and RN_deriv_integral: "integral\<^isup>L N f =
+      (\<integral>x. real (RN_deriv M N x) * f x \<partial>M)" (is ?integral)
 proof -
-  interpret \<nu>: sigma_finite_measure "M\<lparr>measure := \<nu>\<rparr>" by fact
-  have ms: "measure_space (M\<lparr>measure := \<nu>\<rparr>)" by default
+  note ac(2)[simp] and sets_eq_imp_space_eq[OF ac(2), simp]
+  interpret N: sigma_finite_measure N by fact
   have minus_cong: "\<And>A B A' B'::ereal. 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
-  have Nf: "f \<in> borel_measurable (M\<lparr>measure := \<nu>\<rparr>)" using f by simp
+  have Nf: "f \<in> borel_measurable N" using f by (simp add: measurable_def)
   { fix f :: "'a \<Rightarrow> real"
-    { fix x assume *: "RN_deriv M \<nu> x \<noteq> \<infinity>"
-      have "ereal (real (RN_deriv M \<nu> x)) * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)"
+    { fix x assume *: "RN_deriv M N x \<noteq> \<infinity>"
+      have "ereal (real (RN_deriv M N x)) * ereal (f x) = ereal (real (RN_deriv M N x) * f x)"
         by (simp add: mult_le_0_iff)
-      then have "RN_deriv M \<nu> x * ereal (f x) = ereal (real (RN_deriv M \<nu> x) * f x)"
-        using RN_deriv(3)[OF ms \<nu>(2)] * by (auto simp add: ereal_real split: split_if_asm) }
-    then have "(\<integral>\<^isup>+x. ereal (real (RN_deriv M \<nu> x) * f x) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * ereal (f x) \<partial>M)"
-              "(\<integral>\<^isup>+x. ereal (- (real (RN_deriv M \<nu> x) * f x)) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M \<nu> x * ereal (- f x) \<partial>M)"
-      using RN_deriv_finite[OF \<nu>] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric]
+      then have "RN_deriv M N x * ereal (f x) = ereal (real (RN_deriv M N x) * f x)"
+        using RN_deriv(3)[OF ac] * by (auto simp add: ereal_real split: split_if_asm) }
+    then have "(\<integral>\<^isup>+x. ereal (real (RN_deriv M N x) * f x) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M N x * ereal (f x) \<partial>M)"
+              "(\<integral>\<^isup>+x. ereal (- (real (RN_deriv M N x) * f x)) \<partial>M) = (\<integral>\<^isup>+x. RN_deriv M N x * ereal (- f x) \<partial>M)"
+      using RN_deriv_finite[OF N ac] unfolding ereal_mult_minus_right uminus_ereal.simps(1)[symmetric]
       by (auto intro!: positive_integral_cong_AE) }
   note * = this
   show ?integral ?integrable
     unfolding lebesgue_integral_def integrable_def *
-    using f RN_deriv(1)[OF ms \<nu>(2)]
-    by (auto simp: RN_deriv_positive_integral[OF ms \<nu>(2)])
+    using Nf f RN_deriv(1)[OF ac]
+    by (auto simp: RN_deriv_positive_integral[OF ac])
 qed
 
 lemma (in sigma_finite_measure) real_RN_deriv:
-  assumes "finite_measure (M\<lparr>measure := \<nu>\<rparr>)" (is "finite_measure ?\<nu>")
-  assumes ac: "absolutely_continuous \<nu>"
+  assumes "finite_measure N"
+  assumes ac: "absolutely_continuous M N" "sets N = sets M"
   obtains D where "D \<in> borel_measurable M"
-    and "AE x. RN_deriv M \<nu> x = ereal (D x)"
-    and "AE x in M\<lparr>measure := \<nu>\<rparr>. 0 < D x"
+    and "AE x in M. RN_deriv M N x = ereal (D x)"
+    and "AE x in N. 0 < D x"
     and "\<And>x. 0 \<le> D x"
 proof
-  interpret \<nu>: finite_measure ?\<nu> by fact
-  have ms: "measure_space ?\<nu>" by default
-  note RN = RN_deriv[OF ms ac]
+  interpret N: finite_measure N by fact
+  
+  note RN = RN_deriv[OF ac]
 
-  let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M \<nu> x = t}"
+  let ?RN = "\<lambda>t. {x \<in> space M. RN_deriv M N x = t}"
 
-  show "(\<lambda>x. real (RN_deriv M \<nu> x)) \<in> borel_measurable M"
+  show "(\<lambda>x. real (RN_deriv M N x)) \<in> borel_measurable M"
     using RN by auto
 
-  have "\<nu> (?RN \<infinity>) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN \<infinity>) x \<partial>M)"
-    using RN by auto
+  have "N (?RN \<infinity>) = (\<integral>\<^isup>+ x. RN_deriv M N x * indicator (?RN \<infinity>) x \<partial>M)"
+    using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
   also have "\<dots> = (\<integral>\<^isup>+ x. \<infinity> * indicator (?RN \<infinity>) x \<partial>M)"
     by (intro positive_integral_cong) (auto simp: indicator_def)
-  also have "\<dots> = \<infinity> * \<mu> (?RN \<infinity>)"
+  also have "\<dots> = \<infinity> * emeasure M (?RN \<infinity>)"
     using RN by (intro positive_integral_cmult_indicator) auto
-  finally have eq: "\<nu> (?RN \<infinity>) = \<infinity> * \<mu> (?RN \<infinity>)" .
+  finally have eq: "N (?RN \<infinity>) = \<infinity> * emeasure M (?RN \<infinity>)" .
   moreover
-  have "\<mu> (?RN \<infinity>) = 0"
+  have "emeasure M (?RN \<infinity>) = 0"
   proof (rule ccontr)
-    assume "\<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>} \<noteq> 0"
-    moreover from RN have "0 \<le> \<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>}" by auto
-    ultimately have "0 < \<mu> {x \<in> space M. RN_deriv M \<nu> x = \<infinity>}" by auto
-    with eq have "\<nu> (?RN \<infinity>) = \<infinity>" by simp
-    with \<nu>.finite_measure[of "?RN \<infinity>"] RN show False by auto
+    assume "emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>} \<noteq> 0"
+    moreover from RN have "0 \<le> emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}" by auto
+    ultimately have "0 < emeasure M {x \<in> space M. RN_deriv M N x = \<infinity>}" by auto
+    with eq have "N (?RN \<infinity>) = \<infinity>" by simp
+    with N.emeasure_finite[of "?RN \<infinity>"] RN show False by auto
   qed
-  ultimately have "AE x. RN_deriv M \<nu> x < \<infinity>"
+  ultimately have "AE x in M. RN_deriv M N x < \<infinity>"
     using RN by (intro AE_iff_measurable[THEN iffD2]) auto
-  then show "AE x. RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))"
+  then show "AE x in M. RN_deriv M N x = ereal (real (RN_deriv M N x))"
     using RN(3) by (auto simp: ereal_real)
-  then have eq: "AE x in (M\<lparr>measure := \<nu>\<rparr>) . RN_deriv M \<nu> x = ereal (real (RN_deriv M \<nu> x))"
-    using ac absolutely_continuous_AE[OF ms] by auto
+  then have eq: "AE x in N. RN_deriv M N x = ereal (real (RN_deriv M N x))"
+    using ac absolutely_continuous_AE by auto
 
-  show "\<And>x. 0 \<le> real (RN_deriv M \<nu> x)"
+  show "\<And>x. 0 \<le> real (RN_deriv M N x)"
     using RN by (auto intro: real_of_ereal_pos)
 
-  have "\<nu> (?RN 0) = (\<integral>\<^isup>+ x. RN_deriv M \<nu> x * indicator (?RN 0) x \<partial>M)"
-    using RN by auto
+  have "N (?RN 0) = (\<integral>\<^isup>+ x. RN_deriv M N x * indicator (?RN 0) x \<partial>M)"
+    using RN(1,3) by (subst RN(2)[symmetric]) (auto simp: emeasure_density)
   also have "\<dots> = (\<integral>\<^isup>+ x. 0 \<partial>M)"
     by (intro positive_integral_cong) (auto simp: indicator_def)
-  finally have "AE x in (M\<lparr>measure := \<nu>\<rparr>). RN_deriv M \<nu> x \<noteq> 0"
-    using RN by (intro \<nu>.AE_iff_measurable[THEN iffD2]) auto
-  with RN(3) eq show "AE x in (M\<lparr>measure := \<nu>\<rparr>). 0 < real (RN_deriv M \<nu> x)"
+  finally have "AE x in N. RN_deriv M N x \<noteq> 0"
+    using RN by (subst AE_iff_measurable[OF _ refl]) (auto simp: ac cong: sets_eq_imp_space_eq)
+  with RN(3) eq show "AE x in N. 0 < real (RN_deriv M N x)"
     by (auto simp: zero_less_real_of_ereal le_less)
 qed
 
 lemma (in sigma_finite_measure) RN_deriv_singleton:
-  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
-  and ac: "absolutely_continuous \<nu>"
-  and "{x} \<in> sets M"
-  shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}"
+  assumes ac: "absolutely_continuous M N" "sets N = sets M"
+  and x: "{x} \<in> sets M"
+  shows "N {x} = RN_deriv M N x * emeasure M {x}"
 proof -
-  note deriv = RN_deriv[OF assms(1, 2)]
-  from deriv(2)[OF `{x} \<in> sets M`]
-  have "\<nu> {x} = (\<integral>\<^isup>+w. RN_deriv M \<nu> x * indicator {x} w \<partial>M)"
-    by (auto simp: indicator_def intro!: positive_integral_cong)
-  thus ?thesis using positive_integral_cmult_indicator[OF _ `{x} \<in> sets M`] deriv(3)
-    by auto
-qed
-
-theorem (in finite_measure_space) RN_deriv_finite_measure:
-  assumes "measure_space (M\<lparr>measure := \<nu>\<rparr>)"
-  and ac: "absolutely_continuous \<nu>"
-  and "x \<in> space M"
-  shows "\<nu> {x} = RN_deriv M \<nu> x * \<mu> {x}"
-proof -
-  have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto
-  from RN_deriv_singleton[OF assms(1,2) this] show ?thesis .
+  note deriv = RN_deriv[OF ac]
+  from deriv(1,3) `{x} \<in> sets M`
+  have "density M (RN_deriv M N) {x} = (\<integral>\<^isup>+w. RN_deriv M N x * indicator {x} w \<partial>M)"
+    by (auto simp: indicator_def emeasure_density intro!: positive_integral_cong)
+  with x deriv show ?thesis
+    by (auto simp: positive_integral_cmult_indicator)
 qed
 
 end
--- a/src/HOL/Probability/Sigma_Algebra.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/Sigma_Algebra.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -13,6 +13,7 @@
   "~~/src/HOL/Library/Countable"
   "~~/src/HOL/Library/FuncSet"
   "~~/src/HOL/Library/Indicator_Function"
+  "~~/src/HOL/Library/Extended_Real"
 begin
 
 text {* Sigma algebras are an elementary concept in measure
@@ -25,203 +26,198 @@
 
 subsection {* Algebras *}
 
-record 'a algebra =
-  space :: "'a set"
-  sets :: "'a set set"
+locale subset_class =
+  fixes \<Omega> :: "'a set" and M :: "'a set set"
+  assumes space_closed: "M \<subseteq> Pow \<Omega>"
 
-locale subset_class =
-  fixes M :: "('a, 'b) algebra_scheme"
-  assumes space_closed: "sets M \<subseteq> Pow (space M)"
-
-lemma (in subset_class) sets_into_space: "x \<in> sets M \<Longrightarrow> x \<subseteq> space M"
+lemma (in subset_class) sets_into_space: "x \<in> M \<Longrightarrow> x \<subseteq> \<Omega>"
   by (metis PowD contra_subsetD space_closed)
 
 locale ring_of_sets = subset_class +
-  assumes empty_sets [iff]: "{} \<in> sets M"
-     and  Diff [intro]: "\<And>a b. a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a - b \<in> sets M"
-     and  Un [intro]: "\<And>a b. a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<union> b \<in> sets M"
+  assumes empty_sets [iff]: "{} \<in> M"
+     and  Diff [intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a - b \<in> M"
+     and  Un [intro]: "\<And>a b. a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<union> b \<in> M"
 
 lemma (in ring_of_sets) Int [intro]:
-  assumes a: "a \<in> sets M" and b: "b \<in> sets M" shows "a \<inter> b \<in> sets M"
+  assumes a: "a \<in> M" and b: "b \<in> M" shows "a \<inter> b \<in> M"
 proof -
   have "a \<inter> b = a - (a - b)"
     by auto
-  then show "a \<inter> b \<in> sets M"
+  then show "a \<inter> b \<in> M"
     using a b by auto
 qed
 
 lemma (in ring_of_sets) finite_Union [intro]:
-  "finite X \<Longrightarrow> X \<subseteq> sets M \<Longrightarrow> Union X \<in> sets M"
+  "finite X \<Longrightarrow> X \<subseteq> M \<Longrightarrow> Union X \<in> M"
   by (induct set: finite) (auto simp add: Un)
 
 lemma (in ring_of_sets) finite_UN[intro]:
-  assumes "finite I" and "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
-  shows "(\<Union>i\<in>I. A i) \<in> sets M"
+  assumes "finite I" and "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M"
+  shows "(\<Union>i\<in>I. A i) \<in> M"
   using assms by induct auto
 
 lemma (in ring_of_sets) finite_INT[intro]:
-  assumes "finite I" "I \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"
-  shows "(\<Inter>i\<in>I. A i) \<in> sets M"
+  assumes "finite I" "I \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> M"
+  shows "(\<Inter>i\<in>I. A i) \<in> M"
   using assms by (induct rule: finite_ne_induct) auto
 
 lemma (in ring_of_sets) insert_in_sets:
-  assumes "{x} \<in> sets M" "A \<in> sets M" shows "insert x A \<in> sets M"
+  assumes "{x} \<in> M" "A \<in> M" shows "insert x A \<in> M"
 proof -
-  have "{x} \<union> A \<in> sets M" using assms by (rule Un)
+  have "{x} \<union> A \<in> M" using assms by (rule Un)
   thus ?thesis by auto
 qed
 
-lemma (in ring_of_sets) Int_space_eq1 [simp]: "x \<in> sets M \<Longrightarrow> space M \<inter> x = x"
+lemma (in ring_of_sets) Int_space_eq1 [simp]: "x \<in> M \<Longrightarrow> \<Omega> \<inter> x = x"
   by (metis Int_absorb1 sets_into_space)
 
-lemma (in ring_of_sets) Int_space_eq2 [simp]: "x \<in> sets M \<Longrightarrow> x \<inter> space M = x"
+lemma (in ring_of_sets) Int_space_eq2 [simp]: "x \<in> M \<Longrightarrow> x \<inter> \<Omega> = x"
   by (metis Int_absorb2 sets_into_space)
 
 lemma (in ring_of_sets) sets_Collect_conj:
-  assumes "{x\<in>space M. P x} \<in> sets M" "{x\<in>space M. Q x} \<in> sets M"
-  shows "{x\<in>space M. Q x \<and> P x} \<in> sets M"
+  assumes "{x\<in>\<Omega>. P x} \<in> M" "{x\<in>\<Omega>. Q x} \<in> M"
+  shows "{x\<in>\<Omega>. Q x \<and> P x} \<in> M"
 proof -
-  have "{x\<in>space M. Q x \<and> P x} = {x\<in>space M. Q x} \<inter> {x\<in>space M. P x}"
+  have "{x\<in>\<Omega>. Q x \<and> P x} = {x\<in>\<Omega>. Q x} \<inter> {x\<in>\<Omega>. P x}"
     by auto
   with assms show ?thesis by auto
 qed
 
 lemma (in ring_of_sets) sets_Collect_disj:
-  assumes "{x\<in>space M. P x} \<in> sets M" "{x\<in>space M. Q x} \<in> sets M"
-  shows "{x\<in>space M. Q x \<or> P x} \<in> sets M"
+  assumes "{x\<in>\<Omega>. P x} \<in> M" "{x\<in>\<Omega>. Q x} \<in> M"
+  shows "{x\<in>\<Omega>. Q x \<or> P x} \<in> M"
 proof -
-  have "{x\<in>space M. Q x \<or> P x} = {x\<in>space M. Q x} \<union> {x\<in>space M. P x}"
+  have "{x\<in>\<Omega>. Q x \<or> P x} = {x\<in>\<Omega>. Q x} \<union> {x\<in>\<Omega>. P x}"
     by auto
   with assms show ?thesis by auto
 qed
 
 lemma (in ring_of_sets) sets_Collect_finite_All:
-  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>space M. P i x} \<in> sets M" "finite S" "S \<noteq> {}"
-  shows "{x\<in>space M. \<forall>i\<in>S. P i x} \<in> sets M"
+  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S" "S \<noteq> {}"
+  shows "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} \<in> M"
 proof -
-  have "{x\<in>space M. \<forall>i\<in>S. P i x} = (\<Inter>i\<in>S. {x\<in>space M. P i x})"
+  have "{x\<in>\<Omega>. \<forall>i\<in>S. P i x} = (\<Inter>i\<in>S. {x\<in>\<Omega>. P i x})"
     using `S \<noteq> {}` by auto
   with assms show ?thesis by auto
 qed
 
 lemma (in ring_of_sets) sets_Collect_finite_Ex:
-  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>space M. P i x} \<in> sets M" "finite S"
-  shows "{x\<in>space M. \<exists>i\<in>S. P i x} \<in> sets M"
+  assumes "\<And>i. i \<in> S \<Longrightarrow> {x\<in>\<Omega>. P i x} \<in> M" "finite S"
+  shows "{x\<in>\<Omega>. \<exists>i\<in>S. P i x} \<in> M"
 proof -
-  have "{x\<in>space M. \<exists>i\<in>S. P i x} = (\<Union>i\<in>S. {x\<in>space M. P i x})"
+  have "{x\<in>\<Omega>. \<exists>i\<in>S. P i x} = (\<Union>i\<in>S. {x\<in>\<Omega>. P i x})"
     by auto
   with assms show ?thesis by auto
 qed
 
 locale algebra = ring_of_sets +
-  assumes top [iff]: "space M \<in> sets M"
+  assumes top [iff]: "\<Omega> \<in> M"
 
 lemma (in algebra) compl_sets [intro]:
-  "a \<in> sets M \<Longrightarrow> space M - a \<in> sets M"
+  "a \<in> M \<Longrightarrow> \<Omega> - a \<in> M"
   by auto
 
 lemma algebra_iff_Un:
-  "algebra M \<longleftrightarrow>
-    sets M \<subseteq> Pow (space M) &
-    {} \<in> sets M &
-    (\<forall>a \<in> sets M. space M - a \<in> sets M) &
-    (\<forall>a \<in> sets M. \<forall> b \<in> sets M. a \<union> b \<in> sets M)" (is "_ \<longleftrightarrow> ?Un")
+  "algebra \<Omega> M \<longleftrightarrow>
+    M \<subseteq> Pow \<Omega> \<and>
+    {} \<in> M \<and>
+    (\<forall>a \<in> M. \<Omega> - a \<in> M) \<and>
+    (\<forall>a \<in> M. \<forall> b \<in> M. a \<union> b \<in> M)" (is "_ \<longleftrightarrow> ?Un")
 proof
-  assume "algebra M"
-  then interpret algebra M .
+  assume "algebra \<Omega> M"
+  then interpret algebra \<Omega> M .
   show ?Un using sets_into_space by auto
 next
   assume ?Un
-  show "algebra M"
+  show "algebra \<Omega> M"
   proof
-    show space: "sets M \<subseteq> Pow (space M)" "{} \<in> sets M" "space M \<in> sets M"
+    show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M" "\<Omega> \<in> M"
       using `?Un` by auto
-    fix a b assume a: "a \<in> sets M" and b: "b \<in> sets M"
-    then show "a \<union> b \<in> sets M" using `?Un` by auto
-    have "a - b = space M - ((space M - a) \<union> b)"
-      using space a b by auto
-    then show "a - b \<in> sets M"
+    fix a b assume a: "a \<in> M" and b: "b \<in> M"
+    then show "a \<union> b \<in> M" using `?Un` by auto
+    have "a - b = \<Omega> - ((\<Omega> - a) \<union> b)"
+      using \<Omega> a b by auto
+    then show "a - b \<in> M"
       using a b  `?Un` by auto
   qed
 qed
 
 lemma algebra_iff_Int:
-     "algebra M \<longleftrightarrow>
-       sets M \<subseteq> Pow (space M) & {} \<in> sets M &
-       (\<forall>a \<in> sets M. space M - a \<in> sets M) &
-       (\<forall>a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)" (is "_ \<longleftrightarrow> ?Int")
+     "algebra \<Omega> M \<longleftrightarrow>
+       M \<subseteq> Pow \<Omega> & {} \<in> M &
+       (\<forall>a \<in> M. \<Omega> - a \<in> M) &
+       (\<forall>a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)" (is "_ \<longleftrightarrow> ?Int")
 proof
-  assume "algebra M"
-  then interpret algebra M .
+  assume "algebra \<Omega> M"
+  then interpret algebra \<Omega> M .
   show ?Int using sets_into_space by auto
 next
   assume ?Int
-  show "algebra M"
+  show "algebra \<Omega> M"
   proof (unfold algebra_iff_Un, intro conjI ballI)
-    show space: "sets M \<subseteq> Pow (space M)" "{} \<in> sets M"
+    show \<Omega>: "M \<subseteq> Pow \<Omega>" "{} \<in> M"
       using `?Int` by auto
-    from `?Int` show "\<And>a. a \<in> sets M \<Longrightarrow> space M - a \<in> sets M" by auto
-    fix a b assume sets: "a \<in> sets M" "b \<in> sets M"
-    hence "a \<union> b = space M - ((space M - a) \<inter> (space M - b))"
-      using space by blast
-    also have "... \<in> sets M"
-      using sets `?Int` by auto
-    finally show "a \<union> b \<in> sets M" .
+    from `?Int` show "\<And>a. a \<in> M \<Longrightarrow> \<Omega> - a \<in> M" by auto
+    fix a b assume M: "a \<in> M" "b \<in> M"
+    hence "a \<union> b = \<Omega> - ((\<Omega> - a) \<inter> (\<Omega> - b))"
+      using \<Omega> by blast
+    also have "... \<in> M"
+      using M `?Int` by auto
+    finally show "a \<union> b \<in> M" .
   qed
 qed
 
 lemma (in algebra) sets_Collect_neg:
-  assumes "{x\<in>space M. P x} \<in> sets M"
-  shows "{x\<in>space M. \<not> P x} \<in> sets M"
+  assumes "{x\<in>\<Omega>. P x} \<in> M"
+  shows "{x\<in>\<Omega>. \<not> P x} \<in> M"
 proof -
-  have "{x\<in>space M. \<not> P x} = space M - {x\<in>space M. P x}" by auto
+  have "{x\<in>\<Omega>. \<not> P x} = \<Omega> - {x\<in>\<Omega>. P x}" by auto
   with assms show ?thesis by auto
 qed
 
 lemma (in algebra) sets_Collect_imp:
-  "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> {x\<in>space M. Q x} \<in> sets M \<Longrightarrow> {x\<in>space M. Q x \<longrightarrow> P x} \<in> sets M"
+  "{x\<in>\<Omega>. P x} \<in> M \<Longrightarrow> {x\<in>\<Omega>. Q x} \<in> M \<Longrightarrow> {x\<in>\<Omega>. Q x \<longrightarrow> P x} \<in> M"
   unfolding imp_conv_disj by (intro sets_Collect_disj sets_Collect_neg)
 
 lemma (in algebra) sets_Collect_const:
-  "{x\<in>space M. P} \<in> sets M"
+  "{x\<in>\<Omega>. P} \<in> M"
   by (cases P) auto
 
 lemma algebra_single_set:
   assumes "X \<subseteq> S"
-  shows "algebra \<lparr> space = S, sets = { {}, X, S - X, S }\<rparr>"
+  shows "algebra S { {}, X, S - X, S }"
   by default (insert `X \<subseteq> S`, auto)
 
 section {* Restricted algebras *}
 
 abbreviation (in algebra)
-  "restricted_space A \<equiv> \<lparr> space = A, sets = (\<lambda>S. (A \<inter> S)) ` sets M, \<dots> = more M \<rparr>"
+  "restricted_space A \<equiv> (op \<inter> A) ` M"
 
 lemma (in algebra) restricted_algebra:
-  assumes "A \<in> sets M" shows "algebra (restricted_space A)"
+  assumes "A \<in> M" shows "algebra A (restricted_space A)"
   using assms by unfold_locales auto
 
 subsection {* Sigma Algebras *}
 
 locale sigma_algebra = algebra +
-  assumes countable_nat_UN [intro]:
-         "!!A. range A \<subseteq> sets M \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
+  assumes countable_nat_UN [intro]: "\<And>A. range A \<subseteq> M \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
 
 lemma (in algebra) is_sigma_algebra:
-  assumes "finite (sets M)"
-  shows "sigma_algebra M"
+  assumes "finite M"
+  shows "sigma_algebra \<Omega> M"
 proof
-  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M"
-  then have "(\<Union>i. A i) = (\<Union>s\<in>sets M \<inter> range A. s)"
+  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> M"
+  then have "(\<Union>i. A i) = (\<Union>s\<in>M \<inter> range A. s)"
     by auto
-  also have "(\<Union>s\<in>sets M \<inter> range A. s) \<in> sets M"
-    using `finite (sets M)` by auto
-  finally show "(\<Union>i. A i) \<in> sets M" .
+  also have "(\<Union>s\<in>M \<inter> range A. s) \<in> M"
+    using `finite M` by auto
+  finally show "(\<Union>i. A i) \<in> M" .
 qed
 
 lemma countable_UN_eq:
   fixes A :: "'i::countable \<Rightarrow> 'a set"
-  shows "(range A \<subseteq> sets M \<longrightarrow> (\<Union>i. A i) \<in> sets M) \<longleftrightarrow>
-    (range (A \<circ> from_nat) \<subseteq> sets M \<longrightarrow> (\<Union>i. (A \<circ> from_nat) i) \<in> sets M)"
+  shows "(range A \<subseteq> M \<longrightarrow> (\<Union>i. A i) \<in> M) \<longleftrightarrow>
+    (range (A \<circ> from_nat) \<subseteq> M \<longrightarrow> (\<Union>i. (A \<circ> from_nat) i) \<in> M)"
 proof -
   let ?A' = "A \<circ> from_nat"
   have *: "(\<Union>i. ?A' i) = (\<Union>i. A i)" (is "?l = ?r")
@@ -240,60 +236,57 @@
 
 lemma (in sigma_algebra) countable_UN[intro]:
   fixes A :: "'i::countable \<Rightarrow> 'a set"
-  assumes "A`X \<subseteq> sets M"
-  shows  "(\<Union>x\<in>X. A x) \<in> sets M"
+  assumes "A`X \<subseteq> M"
+  shows  "(\<Union>x\<in>X. A x) \<in> M"
 proof -
   let ?A = "\<lambda>i. if i \<in> X then A i else {}"
-  from assms have "range ?A \<subseteq> sets M" by auto
+  from assms have "range ?A \<subseteq> M" by auto
   with countable_nat_UN[of "?A \<circ> from_nat"] countable_UN_eq[of ?A M]
-  have "(\<Union>x. ?A x) \<in> sets M" by auto
+  have "(\<Union>x. ?A x) \<in> M" by auto
   moreover have "(\<Union>x. ?A x) = (\<Union>x\<in>X. A x)" by (auto split: split_if_asm)
   ultimately show ?thesis by simp
 qed
 
 lemma (in sigma_algebra) countable_INT [intro]:
   fixes A :: "'i::countable \<Rightarrow> 'a set"
-  assumes A: "A`X \<subseteq> sets M" "X \<noteq> {}"
-  shows "(\<Inter>i\<in>X. A i) \<in> sets M"
+  assumes A: "A`X \<subseteq> M" "X \<noteq> {}"
+  shows "(\<Inter>i\<in>X. A i) \<in> M"
 proof -
-  from A have "\<forall>i\<in>X. A i \<in> sets M" by fast
-  hence "space M - (\<Union>i\<in>X. space M - A i) \<in> sets M" by blast
+  from A have "\<forall>i\<in>X. A i \<in> M" by fast
+  hence "\<Omega> - (\<Union>i\<in>X. \<Omega> - A i) \<in> M" by blast
   moreover
-  have "(\<Inter>i\<in>X. A i) = space M - (\<Union>i\<in>X. space M - A i)" using space_closed A
+  have "(\<Inter>i\<in>X. A i) = \<Omega> - (\<Union>i\<in>X. \<Omega> - A i)" using space_closed A
     by blast
   ultimately show ?thesis by metis
 qed
 
-lemma ring_of_sets_Pow:
- "ring_of_sets \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
+lemma ring_of_sets_Pow: "ring_of_sets sp (Pow sp)"
   by default auto
 
-lemma algebra_Pow:
-  "algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
+lemma algebra_Pow: "algebra sp (Pow sp)"
   by default auto
 
-lemma sigma_algebra_Pow:
-     "sigma_algebra \<lparr> space = sp, sets = Pow sp, \<dots> = X \<rparr>"
+lemma sigma_algebra_Pow: "sigma_algebra sp (Pow sp)"
   by default auto
 
 lemma sigma_algebra_iff:
-     "sigma_algebra M \<longleftrightarrow>
-      algebra M \<and> (\<forall>A. range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
+  "sigma_algebra \<Omega> M \<longleftrightarrow>
+    algebra \<Omega> M \<and> (\<forall>A. range A \<subseteq> M \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
   by (simp add: sigma_algebra_def sigma_algebra_axioms_def)
 
 lemma (in sigma_algebra) sets_Collect_countable_All:
-  assumes "\<And>i. {x\<in>space M. P i x} \<in> sets M"
-  shows "{x\<in>space M. \<forall>i::'i::countable. P i x} \<in> sets M"
+  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
+  shows "{x\<in>\<Omega>. \<forall>i::'i::countable. P i x} \<in> M"
 proof -
-  have "{x\<in>space M. \<forall>i::'i::countable. P i x} = (\<Inter>i. {x\<in>space M. P i x})" by auto
+  have "{x\<in>\<Omega>. \<forall>i::'i::countable. P i x} = (\<Inter>i. {x\<in>\<Omega>. P i x})" by auto
   with assms show ?thesis by auto
 qed
 
 lemma (in sigma_algebra) sets_Collect_countable_Ex:
-  assumes "\<And>i. {x\<in>space M. P i x} \<in> sets M"
-  shows "{x\<in>space M. \<exists>i::'i::countable. P i x} \<in> sets M"
+  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
+  shows "{x\<in>\<Omega>. \<exists>i::'i::countable. P i x} \<in> M"
 proof -
-  have "{x\<in>space M. \<exists>i::'i::countable. P i x} = (\<Union>i. {x\<in>space M. P i x})" by auto
+  have "{x\<in>\<Omega>. \<exists>i::'i::countable. P i x} = (\<Union>i. {x\<in>\<Omega>. P i x})" by auto
   with assms show ?thesis by auto
 qed
 
@@ -301,9 +294,19 @@
   sets_Collect_imp sets_Collect_disj sets_Collect_conj sets_Collect_neg sets_Collect_const
   sets_Collect_countable_All sets_Collect_countable_Ex sets_Collect_countable_All
 
+lemma (in sigma_algebra) sets_Collect_countable_Ball:
+  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
+  shows "{x\<in>\<Omega>. \<forall>i::'i::countable\<in>X. P i x} \<in> M"
+  unfolding Ball_def by (intro sets_Collect assms)
+
+lemma (in sigma_algebra) sets_Collect_countable_Bex:
+  assumes "\<And>i. {x\<in>\<Omega>. P i x} \<in> M"
+  shows "{x\<in>\<Omega>. \<exists>i::'i::countable\<in>X. P i x} \<in> M"
+  unfolding Bex_def by (intro sets_Collect assms)
+
 lemma sigma_algebra_single_set:
   assumes "X \<subseteq> S"
-  shows "sigma_algebra \<lparr> space = S, sets = { {}, X, S - X, S }\<rparr>"
+  shows "sigma_algebra S { {}, X, S - X, S }"
   using algebra.is_sigma_algebra[OF algebra_single_set[OF `X \<subseteq> S`]] by simp
 
 subsection {* Binary Unions *}
@@ -321,37 +324,34 @@
   by (simp add: INF_def range_binary_eq)
 
 lemma sigma_algebra_iff2:
-     "sigma_algebra M \<longleftrightarrow>
-       sets M \<subseteq> Pow (space M) \<and>
-       {} \<in> sets M \<and> (\<forall>s \<in> sets M. space M - s \<in> sets M) \<and>
-       (\<forall>A. range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
+     "sigma_algebra \<Omega> M \<longleftrightarrow>
+       M \<subseteq> Pow \<Omega> \<and>
+       {} \<in> M \<and> (\<forall>s \<in> M. \<Omega> - s \<in> M) \<and>
+       (\<forall>A. range A \<subseteq> M \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
   by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def
          algebra_iff_Un Un_range_binary)
 
 subsection {* Initial Sigma Algebra *}
 
 text {*Sigma algebras can naturally be created as the closure of any set of
-  sets with regard to the properties just postulated.  *}
+  M with regard to the properties just postulated.  *}
 
 inductive_set
   sigma_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
   for sp :: "'a set" and A :: "'a set set"
   where
-    Basic: "a \<in> A \<Longrightarrow> a \<in> sigma_sets sp A"
+    Basic[intro, simp]: "a \<in> A \<Longrightarrow> a \<in> sigma_sets sp A"
   | Empty: "{} \<in> sigma_sets sp A"
   | Compl: "a \<in> sigma_sets sp A \<Longrightarrow> sp - a \<in> sigma_sets sp A"
   | Union: "(\<And>i::nat. a i \<in> sigma_sets sp A) \<Longrightarrow> (\<Union>i. a i) \<in> sigma_sets sp A"
 
-definition
-  "sigma M = \<lparr> space = space M, sets = sigma_sets (space M) (sets M), \<dots> = more M \<rparr>"
-
 lemma (in sigma_algebra) sigma_sets_subset:
-  assumes a: "a \<subseteq> sets M"
-  shows "sigma_sets (space M) a \<subseteq> sets M"
+  assumes a: "a \<subseteq> M"
+  shows "sigma_sets \<Omega> a \<subseteq> M"
 proof
   fix x
-  assume "x \<in> sigma_sets (space M) a"
-  from this show "x \<in> sets M"
+  assume "x \<in> sigma_sets \<Omega> a"
+  from this show "x \<in> M"
     by (induct rule: sigma_sets.induct, auto) (metis a subsetD)
 qed
 
@@ -359,35 +359,28 @@
   by (erule sigma_sets.induct, auto)
 
 lemma sigma_algebra_sigma_sets:
-     "a \<subseteq> Pow (space M) \<Longrightarrow> sets M = sigma_sets (space M) a \<Longrightarrow> sigma_algebra M"
+     "a \<subseteq> Pow \<Omega> \<Longrightarrow> sigma_algebra \<Omega> (sigma_sets \<Omega> a)"
   by (auto simp add: sigma_algebra_iff2 dest: sigma_sets_into_sp
            intro!: sigma_sets.Union sigma_sets.Empty sigma_sets.Compl)
 
 lemma sigma_sets_least_sigma_algebra:
   assumes "A \<subseteq> Pow S"
-  shows "sigma_sets S A = \<Inter>{B. A \<subseteq> B \<and> sigma_algebra \<lparr>space = S, sets = B\<rparr>}"
+  shows "sigma_sets S A = \<Inter>{B. A \<subseteq> B \<and> sigma_algebra S B}"
 proof safe
-  fix B X assume "A \<subseteq> B" and sa: "sigma_algebra \<lparr> space = S, sets = B \<rparr>"
+  fix B X assume "A \<subseteq> B" and sa: "sigma_algebra S B"
     and X: "X \<in> sigma_sets S A"
   from sigma_algebra.sigma_sets_subset[OF sa, simplified, OF `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"
+  fix X assume "X \<in> \<Inter>{B. A \<subseteq> B \<and> sigma_algebra S B}"
+  then have [intro!]: "\<And>B. A \<subseteq> B \<Longrightarrow> sigma_algebra S B \<Longrightarrow> X \<in> B"
      by simp
-  have "A \<subseteq> sigma_sets S A" using assms
-    by (auto intro!: sigma_sets.Basic)
-  moreover have "sigma_algebra \<lparr>space = S, sets = sigma_sets S A\<rparr>"
+  have "A \<subseteq> sigma_sets S A" using assms by auto
+  moreover have "sigma_algebra S (sigma_sets S A)"
     using assms by (intro sigma_algebra_sigma_sets[of A]) auto
   ultimately show "X \<in> sigma_sets S A" by auto
 qed
 
-lemma sets_sigma: "sets (sigma M) = sigma_sets (space M) (sets M)"
-  unfolding sigma_def by simp
-
-lemma space_sigma [simp]: "space (sigma M) = space M"
-  by (simp add: sigma_def)
-
 lemma sigma_sets_top: "sp \<in> sigma_sets sp A"
   by (metis Diff_empty sigma_sets.Compl sigma_sets.Empty)
 
@@ -421,7 +414,7 @@
   shows "(\<Inter>i\<in>S. a i) \<in> sigma_sets sp A"
 proof -
   from ai have "\<And>i. (if i\<in>S then a i else sp) \<in> sigma_sets sp A"
-    by (simp add: sigma_sets.intros sigma_sets_top)
+    by (simp add: sigma_sets.intros(2-) sigma_sets_top)
   hence "(\<Inter>i. (if i\<in>S then a i else sp)) \<in> sigma_sets sp A"
     by (rule sigma_sets_Inter [OF Asb])
   also have "(\<Inter>i. (if i\<in>S then a i else sp)) = (\<Inter>i\<in>S. a i)"
@@ -430,12 +423,12 @@
 qed
 
 lemma (in sigma_algebra) sigma_sets_eq:
-     "sigma_sets (space M) (sets M) = sets M"
+     "sigma_sets \<Omega> M = M"
 proof
-  show "sets M \<subseteq> sigma_sets (space M) (sets M)"
+  show "M \<subseteq> sigma_sets \<Omega> M"
     by (metis Set.subsetI sigma_sets.Basic)
   next
-  show "sigma_sets (space M) (sets M) \<subseteq> sets M"
+  show "sigma_sets \<Omega> M \<subseteq> M"
     by (metis sigma_sets_subset subset_refl)
 qed
 
@@ -446,52 +439,42 @@
 proof (intro set_eqI iffI)
   fix a assume "a \<in> sigma_sets M A"
   from this A show "a \<in> sigma_sets M B"
-    by induct (auto intro!: sigma_sets.intros del: sigma_sets.Basic)
+    by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)
 next
   fix b assume "b \<in> sigma_sets M B"
   from this B show "b \<in> sigma_sets M A"
-    by induct (auto intro!: sigma_sets.intros del: sigma_sets.Basic)
+    by induct (auto intro!: sigma_sets.intros(2-) del: sigma_sets.Basic)
 qed
 
-lemma sigma_algebra_sigma:
-    "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:
-    "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_subseteq: assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
 proof
   fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
-    by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros)
+    by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros(2-))
 qed
 
 lemma (in sigma_algebra) restriction_in_sets:
   fixes A :: "nat \<Rightarrow> 'a set"
-  assumes "S \<in> sets M"
-  and *: "range A \<subseteq> (\<lambda>A. S \<inter> A) ` sets M" (is "_ \<subseteq> ?r")
-  shows "range A \<subseteq> sets M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` sets M"
+  assumes "S \<in> M"
+  and *: "range A \<subseteq> (\<lambda>A. S \<inter> A) ` M" (is "_ \<subseteq> ?r")
+  shows "range A \<subseteq> M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` M"
 proof -
   { fix i have "A i \<in> ?r" using * by auto
-    hence "\<exists>B. A i = B \<inter> S \<and> B \<in> sets M" by auto
-    hence "A i \<subseteq> S" "A i \<in> sets M" using `S \<in> sets M` by auto }
-  thus "range A \<subseteq> sets M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` sets M"
+    hence "\<exists>B. A i = B \<inter> S \<and> B \<in> M" by auto
+    hence "A i \<subseteq> S" "A i \<in> M" using `S \<in> M` by auto }
+  thus "range A \<subseteq> M" "(\<Union>i. A i) \<in> (\<lambda>A. S \<inter> A) ` M"
     by (auto intro!: image_eqI[of _ _ "(\<Union>i. A i)"])
 qed
 
 lemma (in sigma_algebra) restricted_sigma_algebra:
-  assumes "S \<in> sets M"
-  shows "sigma_algebra (restricted_space S)"
+  assumes "S \<in> M"
+  shows "sigma_algebra S (restricted_space S)"
   unfolding sigma_algebra_def sigma_algebra_axioms_def
 proof safe
-  show "algebra (restricted_space S)" using restricted_algebra[OF assms] .
+  show "algebra S (restricted_space S)" using restricted_algebra[OF assms] .
 next
-  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets (restricted_space S)"
+  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> restricted_space S"
   from restriction_in_sets[OF assms this[simplified]]
-  show "(\<Union>i. A i) \<in> sets (restricted_space S)" by simp
+  show "(\<Union>i. A i) \<in> restricted_space S" by simp
 qed
 
 lemma sigma_sets_Int:
@@ -510,7 +493,7 @@
     then show ?case
       by (auto intro!: sigma_sets.Union
                simp add: UN_extend_simps simp del: UN_simps)
-  qed (auto intro!: sigma_sets.intros)
+  qed (auto intro!: sigma_sets.intros(2-))
   then show "x \<in> sigma_sets A (op \<inter> A ` st)"
     using `A \<subseteq> sp` by (simp add: Int_absorb2)
 next
@@ -529,93 +512,75 @@
     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 (auto intro!: sigma_sets.intros(2-))
 qed
 
-lemma sigma_sets_single[simp]: "sigma_sets {X} {{X}} = {{}, {X}}"
+lemma sigma_sets_empty_eq: "sigma_sets A {} = {{}, A}"
 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
+  fix a assume "a \<in> sigma_sets A {}" then show "a \<in> {{}, A}"
+    by induct blast+
+qed (auto intro: sigma_sets.Empty sigma_sets_top)
+
+lemma sigma_sets_single[simp]: "sigma_sets A {A} = {{}, A}"
+proof (intro set_eqI iffI)
+  fix x assume "x \<in> sigma_sets A {A}"
+  then show "x \<in> {{}, A}"
+    by induct blast+
 next
-  fix x assume "x \<in> {{}, {X}}"
-  then show "x \<in> sigma_sets {X} {{X}}"
+  fix x assume "x \<in> {{}, A}"
+  then show "x \<in> sigma_sets A {A}"
     by (auto intro: sigma_sets.Empty sigma_sets_top)
 qed
 
-lemma (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 in_sigma[intro, simp]: "A \<in> sets M \<Longrightarrow> A \<in> sets (sigma M)"
-  unfolding sigma_def by (auto intro!: sigma_sets.Basic)
-
-lemma (in sigma_algebra) sigma_eq[simp]: "sigma M = M"
-  unfolding sigma_def sigma_sets_eq by simp
-
-lemma sigma_sigma_eq:
-  assumes "sets M \<subseteq> Pow (space M)"
-  shows "sigma (sigma M) = sigma M"
-  using sigma_algebra.sigma_eq[OF sigma_algebra_sigma, OF assms] .
-
 lemma sigma_sets_sigma_sets_eq:
   "M \<subseteq> Pow S \<Longrightarrow> sigma_sets S (sigma_sets S M) = sigma_sets S M"
-  using sigma_sigma_eq[of "\<lparr> space = S, sets = M \<rparr>"]
-  by (simp add: sigma_def)
+  by (rule sigma_algebra.sigma_sets_eq[OF sigma_algebra_sigma_sets, of M S]) auto
 
 lemma sigma_sets_singleton:
   assumes "X \<subseteq> S"
   shows "sigma_sets S { X } = { {}, X, S - X, S }"
 proof -
-  interpret sigma_algebra "\<lparr> space = S, sets = { {}, X, S - X, S }\<rparr>"
+  interpret sigma_algebra S "{ {}, X, S - X, S }"
     by (rule sigma_algebra_single_set) fact
   have "sigma_sets S { X } \<subseteq> sigma_sets S { {}, X, S - X, S }"
     by (rule sigma_sets_subseteq) simp
   moreover have "\<dots> = { {}, X, S - X, S }"
-    using sigma_eq unfolding sigma_def by simp
+    using sigma_sets_eq by simp
   moreover
   { fix A assume "A \<in> { {}, X, S - X, S }"
     then have "A \<in> sigma_sets S { X }"
-      by (auto intro: sigma_sets.intros sigma_sets_top) }
+      by (auto intro: sigma_sets.intros(2-) sigma_sets_top) }
   ultimately have "sigma_sets S { X } = sigma_sets S { {}, X, S - X, S }"
     by (intro antisym) auto
-  with sigma_eq show ?thesis
-    unfolding sigma_def by simp
+  with sigma_sets_eq show ?thesis by simp
 qed
 
 lemma restricted_sigma:
-  assumes S: "S \<in> sets (sigma M)" and M: "sets M \<subseteq> Pow (space M)"
-  shows "algebra.restricted_space (sigma M) S = sigma (algebra.restricted_space M S)"
+  assumes S: "S \<in> sigma_sets \<Omega> M" and M: "M \<subseteq> Pow \<Omega>"
+  shows "algebra.restricted_space (sigma_sets \<Omega> M) S =
+    sigma_sets S (algebra.restricted_space M S)"
 proof -
   from S sigma_sets_into_sp[OF M]
-  have "S \<in> sigma_sets (space M) (sets M)" "S \<subseteq> space M"
-    by (auto simp: sigma_def)
+  have "S \<in> sigma_sets \<Omega> M" "S \<subseteq> \<Omega>" by auto
   from sigma_sets_Int[OF this]
-  show ?thesis
-    by (simp add: sigma_def)
+  show ?thesis by simp
 qed
 
 lemma sigma_sets_vimage_commute:
-  assumes X: "X \<in> space M \<rightarrow> space M'"
-  shows "{X -` A \<inter> space M |A. A \<in> sets (sigma M')}
-       = sigma_sets (space M) {X -` A \<inter> space M |A. A \<in> sets M'}" (is "?L = ?R")
+  assumes X: "X \<in> \<Omega> \<rightarrow> \<Omega>'"
+  shows "{X -` A \<inter> \<Omega> |A. A \<in> sigma_sets \<Omega>' M'}
+       = sigma_sets \<Omega> {X -` A \<inter> \<Omega> |A. A \<in> M'}" (is "?L = ?R")
 proof
   show "?L \<subseteq> ?R"
   proof clarify
-    fix A assume "A \<in> sets (sigma M')"
-    then have "A \<in> sigma_sets (space M') (sets M')" by (simp add: sets_sigma)
-    then show "X -` A \<inter> space M \<in> ?R"
+    fix A assume "A \<in> sigma_sets \<Omega>' M'"
+    then show "X -` A \<inter> \<Omega> \<in> ?R"
     proof induct
-      case (Basic B) then show ?case
-        by (auto intro!: sigma_sets.Basic)
-    next
       case Empty then show ?case
         by (auto intro!: sigma_sets.Empty)
     next
       case (Compl B)
-      have [simp]: "X -` (space M' - B) \<inter> space M = space M - (X -` B \<inter> space M)"
+      have [simp]: "X -` (\<Omega>' - B) \<inter> \<Omega> = \<Omega> - (X -` B \<inter> \<Omega>)"
         by (auto simp add: funcset_mem [OF X])
       with Compl show ?case
         by (auto intro!: sigma_sets.Compl)
@@ -624,194 +589,34 @@
       then show ?case
         by (auto simp add: vimage_UN UN_extend_simps(4) simp del: UN_simps
                  intro!: sigma_sets.Union)
-    qed
+    qed auto
   qed
   show "?R \<subseteq> ?L"
   proof clarify
     fix A assume "A \<in> ?R"
-    then show "\<exists>B. A = X -` B \<inter> space M \<and> B \<in> sets (sigma M')"
+    then show "\<exists>B. A = X -` B \<inter> \<Omega> \<and> B \<in> sigma_sets \<Omega>' M'"
     proof induct
       case (Basic B) then show ?case by auto
     next
       case Empty then show ?case
-        by (auto simp: sets_sigma intro!: sigma_sets.Empty exI[of _ "{}"])
+        by (auto intro!: sigma_sets.Empty exI[of _ "{}"])
     next
       case (Compl B)
-      then obtain A where A: "B = X -` A \<inter> space M" "A \<in> sets (sigma M')" by auto
-      then have [simp]: "space M - B = X -` (space M' - A) \<inter> space M"
+      then obtain A where A: "B = X -` A \<inter> \<Omega>" "A \<in> sigma_sets \<Omega>' M'" by auto
+      then have [simp]: "\<Omega> - B = X -` (\<Omega>' - A) \<inter> \<Omega>"
         by (auto simp add: funcset_mem [OF X])
       with A(2) show ?case
-        by (auto simp: sets_sigma intro: sigma_sets.Compl)
+        by (auto intro: sigma_sets.Compl)
     next
       case (Union F)
-      then have "\<forall>i. \<exists>B. F i = X -` B \<inter> space M \<and> B \<in> sets (sigma M')" by auto
+      then have "\<forall>i. \<exists>B. F i = X -` B \<inter> \<Omega> \<and> B \<in> sigma_sets \<Omega>' M'" by auto
       from choice[OF this] guess A .. note A = this
       with A show ?case
-        by (auto simp: sets_sigma vimage_UN[symmetric] intro: sigma_sets.Union)
+        by (auto simp: vimage_UN[symmetric] intro: sigma_sets.Union)
     qed
   qed
 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: "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 (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 (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
-            by (auto simp add: ba) (metis B subsetD PowD)
-        next
-          case Empty
-          thus ?case
-            by auto
-        next
-          case (Compl a)
-          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
-        qed
-    qed
-  thus ?thesis
-    by (simp add: measurable_def sigma_algebra_axioms sigma_algebra_sigma B f)
-       (auto simp add: sigma_def)
-qed
-
-lemma measurable_cong:
-  assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
-  shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
-  unfolding measurable_def using assms
-  by (simp cong: vimage_inter_cong Pi_cong)
-
-lemma measurable_space:
-  "f \<in> measurable M A \<Longrightarrow> x \<in> space M \<Longrightarrow> f x \<in> space A"
-   unfolding measurable_def by auto
-
-lemma measurable_sets:
-  "f \<in> measurable M A \<Longrightarrow> S \<in> sets A \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
-   unfolding measurable_def by auto
-
-lemma (in sigma_algebra) measurable_subset:
-     "(\<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:
-     "\<lbrakk> space m1 = space m1' ; space m2 = space m2' ;
-        sets m1 = sets m1' ; sets m2 = sets m2' \<rbrakk>
-      \<Longrightarrow> measurable m1 m2 = measurable m1' m2'"
-  by (simp add: measurable_def sigma_algebra_iff2)
-
-lemma (in sigma_algebra) measurable_const[intro, simp]:
-  assumes "c \<in> space M'"
-  shows "(\<lambda>x. c) \<in> measurable M M'"
-  using assms by (auto simp add: measurable_def)
-
-lemma (in sigma_algebra) measurable_If:
-  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
-  assumes P: "{x\<in>space M. P x} \<in> sets M"
-  shows "(\<lambda>x. if P x then f x else g x) \<in> measurable M M'"
-  unfolding measurable_def
-proof safe
-  fix x assume "x \<in> space M"
-  thus "(if P x then f x else g x) \<in> space M'"
-    using measure unfolding measurable_def by auto
-next
-  fix A assume "A \<in> sets M'"
-  hence *: "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M =
-    ((f -` A \<inter> space M) \<inter> {x\<in>space M. P x}) \<union>
-    ((g -` A \<inter> space M) \<inter> (space M - {x\<in>space M. P x}))"
-    using measure unfolding measurable_def by (auto split: split_if_asm)
-  show "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M \<in> sets M"
-    using `A \<in> sets M'` measure P unfolding * measurable_def
-    by (auto intro!: Un)
-qed
-
-lemma (in sigma_algebra) measurable_If_set:
-  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
-  assumes P: "A \<in> sets M"
-  shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"
-proof (rule measurable_If[OF measure])
-  have "{x \<in> space M. x \<in> A} = A" using `A \<in> sets M` sets_into_space by auto
-  thus "{x \<in> space M. x \<in> A} \<in> sets M" using `A \<in> sets M` by auto
-qed
-
-lemma (in ring_of_sets) measurable_ident[intro, simp]: "id \<in> measurable M M"
-  by (auto simp add: measurable_def)
-
-lemma measurable_comp[intro]:
-  fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
-  shows "f \<in> measurable a b \<Longrightarrow> g \<in> measurable b c \<Longrightarrow> (g o f) \<in> measurable a c"
-  apply (auto simp add: measurable_def vimage_compose)
-  apply (subgoal_tac "f -` g -` y \<inter> space a = f -` (g -` y \<inter> space b) \<inter> space a")
-  apply force+
-  done
-
-lemma measurable_strong:
-  fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
-  assumes f: "f \<in> measurable a b" and g: "g \<in> (space b -> space c)"
-      and a: "sigma_algebra a" and b: "sigma_algebra b" and c: "sigma_algebra c"
-      and t: "f ` (space a) \<subseteq> t"
-      and cb: "\<And>s. s \<in> sets c \<Longrightarrow> (g -` s) \<inter> t \<in> sets b"
-  shows "(g o f) \<in> measurable a c"
-proof -
-  have fab: "f \<in> (space a -> space b)"
-   and ba: "\<And>y. y \<in> sets b \<Longrightarrow> (f -` y) \<inter> (space a) \<in> sets a" using f
-     by (auto simp add: measurable_def)
-  have eq: "f -` g -` y \<inter> space a = f -` (g -` y \<inter> t) \<inter> space a" using t
-    by force
-  show ?thesis
-    apply (auto simp add: measurable_def vimage_compose a c)
-    apply (metis funcset_mem fab g)
-    apply (subst eq, metis ba cb)
-    done
-qed
-
-lemma measurable_mono1:
-     "a \<subseteq> b \<Longrightarrow> sigma_algebra \<lparr>space = X, sets = b\<rparr>
-      \<Longrightarrow> measurable \<lparr>space = X, sets = a\<rparr> c \<subseteq> measurable \<lparr>space = X, sets = b\<rparr> c"
-  by (auto simp add: measurable_def)
-
-lemma measurable_up_sigma:
-  "measurable A M \<subseteq> measurable (sigma A) M"
-  unfolding measurable_def
-  by (auto simp: sigma_def intro: sigma_sets.Basic)
-
-lemma (in sigma_algebra) measurable_range_reduce:
-   "\<lbrakk> f \<in> measurable M \<lparr>space = s, sets = Pow s\<rparr> ; s \<noteq> {} \<rbrakk>
-    \<Longrightarrow> f \<in> measurable M \<lparr>space = s \<inter> (f ` space M), sets = Pow (s \<inter> (f ` space M))\<rparr>"
-  by (simp add: measurable_def sigma_algebra_Pow del: Pow_Int_eq) blast
-
-lemma (in sigma_algebra) measurable_Pow_to_Pow:
-   "(sets M = Pow (space M)) \<Longrightarrow> f \<in> measurable M \<lparr>space = UNIV, sets = Pow UNIV\<rparr>"
-  by (auto simp add: measurable_def sigma_algebra_def sigma_algebra_axioms_def algebra_def)
-
-lemma (in sigma_algebra) measurable_Pow_to_Pow_image:
-   "sets M = Pow (space M)
-    \<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) 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 (fastforce simp: measurable_def sets_sigma intro: sigma_sets.intros)
-
 section "Disjoint families"
 
 definition
@@ -906,8 +711,8 @@
 
 lemma (in ring_of_sets) UNION_in_sets:
   fixes A:: "nat \<Rightarrow> 'a set"
-  assumes A: "range A \<subseteq> sets M "
-  shows  "(\<Union>i\<in>{0..<n}. A i) \<in> sets M"
+  assumes A: "range A \<subseteq> M"
+  shows  "(\<Union>i\<in>{0..<n}. A i) \<in> M"
 proof (induct n)
   case 0 show ?case by simp
 next
@@ -917,16 +722,16 @@
 qed
 
 lemma (in ring_of_sets) range_disjointed_sets:
-  assumes A: "range A \<subseteq> sets M "
-  shows  "range (disjointed A) \<subseteq> sets M"
+  assumes A: "range A \<subseteq> M"
+  shows  "range (disjointed A) \<subseteq> M"
 proof (auto simp add: disjointed_def)
   fix n
-  show "A n - (\<Union>i\<in>{0..<n}. A i) \<in> sets M" using UNION_in_sets
+  show "A n - (\<Union>i\<in>{0..<n}. A i) \<in> M" using UNION_in_sets
     by (metis A Diff UNIV_I image_subset_iff)
 qed
 
 lemma (in algebra) range_disjointed_sets':
-  "range A \<subseteq> sets M \<Longrightarrow> range (disjointed A) \<subseteq> sets M"
+  "range A \<subseteq> M \<Longrightarrow> range (disjointed A) \<subseteq> M"
   using range_disjointed_sets .
 
 lemma disjointed_0[simp]: "disjointed A 0 = A 0"
@@ -942,81 +747,518 @@
   by (simp add: disjointed_def atLeastLessThanSuc_atLeastAtMost atLeast0AtMost)
 
 lemma sigma_algebra_disjoint_iff:
-     "sigma_algebra M \<longleftrightarrow>
-      algebra M &
-      (\<forall>A. range A \<subseteq> sets M \<longrightarrow> disjoint_family A \<longrightarrow>
-           (\<Union>i::nat. A i) \<in> sets M)"
+  "sigma_algebra \<Omega> M \<longleftrightarrow> algebra \<Omega> M \<and>
+    (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
 proof (auto simp add: sigma_algebra_iff)
   fix A :: "nat \<Rightarrow> 'a set"
-  assume M: "algebra M"
-     and A: "range A \<subseteq> sets M"
-     and UnA: "\<forall>A. range A \<subseteq> sets M \<longrightarrow>
-               disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
-  hence "range (disjointed A) \<subseteq> sets M \<longrightarrow>
+  assume M: "algebra \<Omega> M"
+     and A: "range A \<subseteq> M"
+     and UnA: "\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M"
+  hence "range (disjointed A) \<subseteq> M \<longrightarrow>
          disjoint_family (disjointed A) \<longrightarrow>
-         (\<Union>i. disjointed A i) \<in> sets M" by blast
-  hence "(\<Union>i. disjointed A i) \<in> sets M"
-    by (simp add: algebra.range_disjointed_sets' M A disjoint_family_disjointed)
-  thus "(\<Union>i::nat. A i) \<in> sets M" by (simp add: UN_disjointed_eq)
+         (\<Union>i. disjointed A i) \<in> M" by blast
+  hence "(\<Union>i. disjointed A i) \<in> M"
+    by (simp add: algebra.range_disjointed_sets'[of \<Omega>] M A disjoint_family_disjointed)
+  thus "(\<Union>i::nat. A i) \<in> M" by (simp add: UN_disjointed_eq)
+qed
+
+section {* Measure type *}
+
+definition positive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
+  "positive M \<mu> \<longleftrightarrow> \<mu> {} = 0 \<and> (\<forall>A\<in>M. 0 \<le> \<mu> A)"
+
+definition countably_additive :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
+  "countably_additive M f \<longleftrightarrow> (\<forall>A. range A \<subseteq> M \<longrightarrow> disjoint_family A \<longrightarrow> (\<Union>i. A i) \<in> M \<longrightarrow>
+    (\<Sum>i. f (A i)) = f (\<Union>i. A i))"
+
+definition measure_space :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
+  "measure_space \<Omega> A \<mu> \<longleftrightarrow> sigma_algebra \<Omega> A \<and> positive A \<mu> \<and> countably_additive A \<mu>"
+
+typedef (open) 'a measure = "{(\<Omega>::'a set, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu> }"
+proof
+  have "sigma_algebra UNIV {{}, UNIV}"
+    proof qed auto
+  then show "(UNIV, {{}, UNIV}, \<lambda>A. 0) \<in> {(\<Omega>, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu>} "
+    by (auto simp: measure_space_def positive_def countably_additive_def)
+qed
+
+definition space :: "'a measure \<Rightarrow> 'a set" where
+  "space M = fst (Rep_measure M)"
+
+definition sets :: "'a measure \<Rightarrow> 'a set set" where
+  "sets M = fst (snd (Rep_measure M))"
+
+definition emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ereal" where
+  "emeasure M = snd (snd (Rep_measure M))"
+
+definition measure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> real" where
+  "measure M A = real (emeasure M A)"
+
+declare [[coercion sets]]
+
+declare [[coercion measure]]
+
+declare [[coercion emeasure]]
+
+lemma measure_space: "measure_space (space M) (sets M) (emeasure M)"
+  by (cases M) (auto simp: space_def sets_def emeasure_def Abs_measure_inverse)
+
+interpretation sigma_algebra "space M" "sets M" for M :: "'a measure"
+  using measure_space[of M] by (auto simp: measure_space_def)
+
+definition measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
+  "measure_of \<Omega> A \<mu> = Abs_measure (\<Omega>, sigma_sets \<Omega> A,
+    \<lambda>a. if a \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> a else 0)"
+
+abbreviation "sigma \<Omega> A \<equiv> measure_of \<Omega> A (\<lambda>x. 0)"
+
+lemma measure_space_0: "A \<subseteq> Pow \<Omega> \<Longrightarrow> measure_space \<Omega> (sigma_sets \<Omega> A) (\<lambda>x. 0)"
+  unfolding measure_space_def
+  by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def)
+
+lemma (in ring_of_sets) positive_cong_eq:
+  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> positive M \<mu>' = positive M \<mu>"
+  by (auto simp add: positive_def)
+
+lemma (in sigma_algebra) countably_additive_eq:
+  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> countably_additive M \<mu>' = countably_additive M \<mu>"
+  unfolding countably_additive_def
+  by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq)
+
+lemma measure_space_eq:
+  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a"
+  shows "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
+proof -
+  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" using closed by (rule sigma_algebra_sigma_sets)
+  from positive_cong_eq[OF eq, of "\<lambda>i. i"] countably_additive_eq[OF eq, of "\<lambda>i. i"] show ?thesis
+    by (auto simp: measure_space_def)
+qed
+
+lemma measure_of_eq:
+  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "(\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a)"
+  shows "measure_of \<Omega> A \<mu> = measure_of \<Omega> A \<mu>'"
+proof -
+  have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
+    using assms by (rule measure_space_eq)
+  with eq show ?thesis
+    by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure])
+qed
+
+lemma
+  assumes A: "A \<subseteq> Pow \<Omega>"
+  shows sets_measure_of[simp]: "sets (measure_of \<Omega> A \<mu>) = sigma_sets \<Omega> A" (is ?sets)
+    and space_measure_of[simp]: "space (measure_of \<Omega> A \<mu>) = \<Omega>" (is ?space)
+proof -
+  have "?sets \<and> ?space"
+  proof cases
+    assume "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>"
+    moreover have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A)
+       (\<lambda>a. if a \<in> sigma_sets \<Omega> A then \<mu> a else 0)"
+      using A by (rule measure_space_eq) auto
+    ultimately show "?sets \<and> ?space"
+      by (auto simp: Abs_measure_inverse measure_of_def sets_def space_def)
+  next
+    assume "\<not> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>"
+    with A show "?sets \<and> ?space"
+      by (auto simp: Abs_measure_inverse measure_of_def sets_def space_def measure_space_0)
+  qed
+  then show ?sets ?space by auto
+qed
+
+lemma (in sigma_algebra) sets_measure_of_eq[simp]:
+  "sets (measure_of \<Omega> M \<mu>) = M"
+  using space_closed by (auto intro!: sigma_sets_eq)
+
+lemma (in sigma_algebra) space_measure_of_eq[simp]:
+  "space (measure_of \<Omega> M \<mu>) = \<Omega>"
+  using space_closed by (auto intro!: sigma_sets_eq)
+
+lemma measure_of_subset:
+  "M \<subseteq> Pow \<Omega> \<Longrightarrow> M' \<subseteq> M \<Longrightarrow> sets (measure_of \<Omega> M' \<mu>) \<subseteq> sets (measure_of \<Omega> M \<mu>')"
+  by (auto intro!: sigma_sets_subseteq)
+
+lemma in_extended_measure[intro, simp]: "M \<subseteq> Pow \<Omega> \<Longrightarrow> A \<in> M \<Longrightarrow> A \<in> sets (measure_of \<Omega> M \<mu>)"
+  by auto
+
+section {* Constructing simple @{typ "'a measure"} *}
+
+lemma emeasure_measure_of:
+  assumes M: "M = measure_of \<Omega> A \<mu>"
+  assumes ms: "A \<subseteq> Pow \<Omega>" "positive (sets M) \<mu>" "countably_additive (sets M) \<mu>"
+  assumes X: "X \<in> sets M"
+  shows "emeasure M X = \<mu> X"
+proof -
+  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" by (rule sigma_algebra_sigma_sets) fact
+  have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>"
+    using ms M by (simp add: measure_space_def sigma_algebra_sigma_sets)
+  moreover have "measure_space \<Omega> (sigma_sets \<Omega> A) (\<lambda>a. if a \<in> sigma_sets \<Omega> A then \<mu> a else 0)
+    = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>"
+    using ms(1) by (rule measure_space_eq) auto
+  moreover have "X \<in> sigma_sets \<Omega> A"
+    using X M ms by simp
+  ultimately show ?thesis
+    unfolding emeasure_def measure_of_def M
+    by (subst Abs_measure_inverse) (simp_all add: sigma_sets_eq)
+qed
+
+lemma emeasure_measure_of_sigma:
+  assumes ms: "sigma_algebra \<Omega> M" "positive M \<mu>" "countably_additive M \<mu>"
+  assumes A: "A \<in> M"
+  shows "emeasure (measure_of \<Omega> M \<mu>) A = \<mu> A"
+proof -
+  interpret sigma_algebra \<Omega> M by fact
+  have "measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
+    using ms sigma_sets_eq by (simp add: measure_space_def)
+  moreover have "measure_space \<Omega> (sigma_sets \<Omega> M) (\<lambda>a. if a \<in> sigma_sets \<Omega> M then \<mu> a else 0)
+    = measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
+    using space_closed by (rule measure_space_eq) auto
+  ultimately show ?thesis using A
+    unfolding emeasure_def measure_of_def
+    by (subst Abs_measure_inverse) (simp_all add: sigma_sets_eq)
+qed
+
+lemma measure_cases[cases type: measure]:
+  obtains (measure) \<Omega> A \<mu> where "x = Abs_measure (\<Omega>, A, \<mu>)" "\<forall>a\<in>-A. \<mu> a = 0" "measure_space \<Omega> A \<mu>"
+  by atomize_elim (cases x, auto)
+
+lemma sets_eq_imp_space_eq:
+  "sets M = sets M' \<Longrightarrow> space M = space M'"
+  using top[of M] top[of M'] space_closed[of M] space_closed[of M']
+  by blast
+
+lemma emeasure_notin_sets: "A \<notin> sets M \<Longrightarrow> emeasure M A = 0"
+  by (cases M) (auto simp: sets_def emeasure_def Abs_measure_inverse measure_space_def)
+
+lemma measure_notin_sets: "A \<notin> sets M \<Longrightarrow> measure M A = 0"
+  by (simp add: measure_def emeasure_notin_sets)
+
+lemma measure_eqI:
+  fixes M N :: "'a measure"
+  assumes "sets M = sets N" and eq: "\<And>A. A \<in> sets M \<Longrightarrow> emeasure M A = emeasure N A"
+  shows "M = N"
+proof (cases M N rule: measure_cases[case_product measure_cases])
+  case (measure_measure \<Omega> A \<mu> \<Omega>' A' \<mu>')
+  interpret M: sigma_algebra \<Omega> A using measure_measure by (auto simp: measure_space_def)
+  interpret N: sigma_algebra \<Omega>' A' using measure_measure by (auto simp: measure_space_def)
+  have "A = sets M" "A' = sets N"
+    using measure_measure by (simp_all add: sets_def Abs_measure_inverse)
+  with `sets M = sets N` have "A = A'" by simp
+  moreover with M.top N.top M.space_closed N.space_closed have "\<Omega> = \<Omega>'" by auto
+  moreover { fix B have "\<mu> B = \<mu>' B"
+    proof cases
+      assume "B \<in> A"
+      with eq `A = sets M` have "emeasure M B = emeasure N B" by simp
+      with measure_measure show "\<mu> B = \<mu>' B"
+        by (simp add: emeasure_def Abs_measure_inverse)
+    next
+      assume "B \<notin> A"
+      with `A = sets M` `A' = sets N` `A = A'` have "B \<notin> sets M" "B \<notin> sets N"
+        by auto
+      then have "emeasure M B = 0" "emeasure N B = 0"
+        by (simp_all add: emeasure_notin_sets)
+      with measure_measure show "\<mu> B = \<mu>' B"
+        by (simp add: emeasure_def Abs_measure_inverse)
+    qed }
+  then have "\<mu> = \<mu>'" by auto
+  ultimately show "M = N"
+    by (simp add: measure_measure)
 qed
 
+lemma emeasure_sigma: "A \<subseteq> Pow \<Omega> \<Longrightarrow> emeasure (sigma \<Omega> A) = (\<lambda>_. 0)"
+  using measure_space_0[of A \<Omega>]
+  by (simp add: measure_of_def emeasure_def Abs_measure_inverse)
+
+lemma sigma_eqI:
+  assumes [simp]: "M \<subseteq> Pow \<Omega>" "N \<subseteq> Pow \<Omega>" "sigma_sets \<Omega> M = sigma_sets \<Omega> N"
+  shows "sigma \<Omega> M = sigma \<Omega> N"
+  by (rule measure_eqI) (simp_all add: emeasure_sigma)
+
+section {* Measurable functions *}
+
+definition measurable :: "'a measure \<Rightarrow> 'b measure \<Rightarrow> ('a \<Rightarrow> 'b) set" where
+  "measurable A B = {f \<in> space A -> space B. \<forall>y \<in> sets B. f -` y \<inter> space A \<in> sets A}"
+
+lemma measurable_space:
+  "f \<in> measurable M A \<Longrightarrow> x \<in> space M \<Longrightarrow> f x \<in> space A"
+   unfolding measurable_def by auto
+
+lemma measurable_sets:
+  "f \<in> measurable M A \<Longrightarrow> S \<in> sets A \<Longrightarrow> f -` S \<inter> space M \<in> sets M"
+   unfolding measurable_def by auto
+
+lemma measurable_sigma_sets:
+  assumes B: "sets N = sigma_sets \<Omega> A" "A \<subseteq> Pow \<Omega>"
+      and f: "f \<in> space M \<rightarrow> \<Omega>"
+      and ba: "\<And>y. y \<in> A \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
+  shows "f \<in> measurable M N"
+proof -
+  interpret A: sigma_algebra \<Omega> "sigma_sets \<Omega> A" using B(2) by (rule sigma_algebra_sigma_sets)
+  from B top[of N] A.top space_closed[of N] A.space_closed have \<Omega>: "\<Omega> = space N" by force
+  
+  { fix X assume "X \<in> sigma_sets \<Omega> A"
+    then have "f -` X \<inter> space M \<in> sets M \<and> X \<subseteq> \<Omega>"
+      proof induct
+        case (Basic a) then show ?case
+          by (auto simp add: ba) (metis B(2) subsetD PowD)
+      next
+        case (Compl a)
+        have [simp]: "f -` \<Omega> \<inter> space M = space M"
+          by (auto simp add: funcset_mem [OF f])
+        then show ?case
+          by (auto simp add: vimage_Diff Diff_Int_distrib2 compl_sets Compl)
+      next
+        case (Union a)
+        then show ?case
+          by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast
+      qed auto }
+  with f show ?thesis
+    by (auto simp add: measurable_def B \<Omega>)
+qed
+
+lemma measurable_measure_of:
+  assumes B: "N \<subseteq> Pow \<Omega>"
+      and f: "f \<in> space M \<rightarrow> \<Omega>"
+      and ba: "\<And>y. y \<in> N \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
+  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>)"
+proof -
+  have "sets (measure_of \<Omega> N \<mu>) = sigma_sets \<Omega> N"
+    using B by (rule sets_measure_of)
+  from this assms show ?thesis by (rule measurable_sigma_sets)
+qed
+
+lemma measurable_iff_measure_of:
+  assumes "N \<subseteq> Pow \<Omega>" "f \<in> space M \<rightarrow> \<Omega>"
+  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>) \<longleftrightarrow> (\<forall>A\<in>N. f -` A \<inter> space M \<in> sets M)"
+  by (metis assms in_extended_measure measurable_measure_of assms measurable_sets)
+
+lemma measurable_cong:
+  assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
+  shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
+  unfolding measurable_def using assms
+  by (simp cong: vimage_inter_cong Pi_cong)
+
+lemma measurable_eqI:
+     "\<lbrakk> space m1 = space m1' ; space m2 = space m2' ;
+        sets m1 = sets m1' ; sets m2 = sets m2' \<rbrakk>
+      \<Longrightarrow> measurable m1 m2 = measurable m1' m2'"
+  by (simp add: measurable_def sigma_algebra_iff2)
+
+lemma measurable_const[intro, simp]:
+  "c \<in> space M' \<Longrightarrow> (\<lambda>x. c) \<in> measurable M M'"
+  by (auto simp add: measurable_def)
+
+lemma measurable_If:
+  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
+  assumes P: "{x\<in>space M. P x} \<in> sets M"
+  shows "(\<lambda>x. if P x then f x else g x) \<in> measurable M M'"
+  unfolding measurable_def
+proof safe
+  fix x assume "x \<in> space M"
+  thus "(if P x then f x else g x) \<in> space M'"
+    using measure unfolding measurable_def by auto
+next
+  fix A assume "A \<in> sets M'"
+  hence *: "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M =
+    ((f -` A \<inter> space M) \<inter> {x\<in>space M. P x}) \<union>
+    ((g -` A \<inter> space M) \<inter> (space M - {x\<in>space M. P x}))"
+    using measure unfolding measurable_def by (auto split: split_if_asm)
+  show "(\<lambda>x. if P x then f x else g x) -` A \<inter> space M \<in> sets M"
+    using `A \<in> sets M'` measure P unfolding * measurable_def
+    by (auto intro!: Un)
+qed
+
+lemma measurable_If_set:
+  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
+  assumes P: "A \<in> sets M"
+  shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"
+proof (rule measurable_If[OF measure])
+  have "{x \<in> space M. x \<in> A} = A" using `A \<in> sets M` sets_into_space by auto
+  thus "{x \<in> space M. x \<in> A} \<in> sets M" using `A \<in> sets M` by auto
+qed
+
+lemma measurable_ident[intro, simp]: "id \<in> measurable M M"
+  by (auto simp add: measurable_def)
+
+lemma measurable_comp[intro]:
+  fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
+  shows "f \<in> measurable a b \<Longrightarrow> g \<in> measurable b c \<Longrightarrow> (g o f) \<in> measurable a c"
+  apply (auto simp add: measurable_def vimage_compose)
+  apply (subgoal_tac "f -` g -` y \<inter> space a = f -` (g -` y \<inter> space b) \<inter> space a")
+  apply force+
+  done
+
+lemma measurable_Least:
+  assumes meas: "\<And>i::nat. {x\<in>space M. P i x} \<in> M"
+  shows "(\<lambda>x. LEAST j. P j x) -` A \<inter> space M \<in> sets M"
+proof -
+  { fix i have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M \<in> sets M"
+    proof cases
+      assume i: "(LEAST j. False) = i"
+      have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
+        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x})) \<union> (space M - (\<Union>i. {x\<in>space M. P i x}))"
+        by (simp add: set_eq_iff, safe)
+           (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality)
+      with meas show ?thesis
+        by (auto intro!: Int)
+    next
+      assume i: "(LEAST j. False) \<noteq> i"
+      then have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
+        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x}))"
+      proof (simp add: set_eq_iff, safe)
+        fix x assume neq: "(LEAST j. False) \<noteq> (LEAST j. P j x)"
+        have "\<exists>j. P j x"
+          by (rule ccontr) (insert neq, auto)
+        then show "P (LEAST j. P j x) x" by (rule LeastI_ex)
+      qed (auto dest: Least_le intro!: Least_equality)
+      with meas show ?thesis
+        by auto
+    qed }
+  then have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) \<in> sets M"
+    by (intro countable_UN) auto
+  moreover have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) =
+    (\<lambda>x. LEAST j. P j x) -` A \<inter> space M" by auto
+  ultimately show ?thesis by auto
+qed
+
+lemma measurable_strong:
+  fixes f :: "'a \<Rightarrow> 'b" and g :: "'b \<Rightarrow> 'c"
+  assumes f: "f \<in> measurable a b" and g: "g \<in> space b \<rightarrow> space c"
+      and t: "f ` (space a) \<subseteq> t"
+      and cb: "\<And>s. s \<in> sets c \<Longrightarrow> (g -` s) \<inter> t \<in> sets b"
+  shows "(g o f) \<in> measurable a c"
+proof -
+  have fab: "f \<in> (space a -> space b)"
+   and ba: "\<And>y. y \<in> sets b \<Longrightarrow> (f -` y) \<inter> (space a) \<in> sets a" using f
+     by (auto simp add: measurable_def)
+  have eq: "\<And>y. f -` g -` y \<inter> space a = f -` (g -` y \<inter> t) \<inter> space a" using t
+    by force
+  show ?thesis
+    apply (auto simp add: measurable_def vimage_compose)
+    apply (metis funcset_mem fab g)
+    apply (subst eq, metis ba cb)
+    done
+qed
+
+lemma measurable_mono1:
+  "M' \<subseteq> Pow \<Omega> \<Longrightarrow> M \<subseteq> M' \<Longrightarrow>
+    measurable (measure_of \<Omega> M \<mu>) N \<subseteq> measurable (measure_of \<Omega> M' \<mu>') N"
+  using measure_of_subset[of M' \<Omega> M] by (auto simp add: measurable_def)
+
+subsection {* Extend measure *}
+
+definition "extend_measure \<Omega> I G \<mu> =
+  (if (\<exists>\<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>') \<and> \<not> (\<forall>i\<in>I. \<mu> i = 0)
+      then measure_of \<Omega> (G`I) (SOME \<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>')
+      else measure_of \<Omega> (G`I) (\<lambda>_. 0))"
+
+lemma space_extend_measure: "G ` I \<subseteq> Pow \<Omega> \<Longrightarrow> space (extend_measure \<Omega> I G \<mu>) = \<Omega>"
+  unfolding extend_measure_def by simp
+
+lemma sets_extend_measure: "G ` I \<subseteq> Pow \<Omega> \<Longrightarrow> sets (extend_measure \<Omega> I G \<mu>) = sigma_sets \<Omega> (G`I)"
+  unfolding extend_measure_def by simp
+
+lemma emeasure_extend_measure:
+  assumes M: "M = extend_measure \<Omega> I G \<mu>"
+    and eq: "\<And>i. i \<in> I \<Longrightarrow> \<mu>' (G i) = \<mu> i"
+    and ms: "G ` I \<subseteq> Pow \<Omega>" "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"
+    and "i \<in> I"
+  shows "emeasure M (G i) = \<mu> i"
+proof cases
+  assume *: "(\<forall>i\<in>I. \<mu> i = 0)"
+  with M have M_eq: "M = measure_of \<Omega> (G`I) (\<lambda>_. 0)"
+   by (simp add: extend_measure_def)
+  from measure_space_0[OF ms(1)] ms `i\<in>I`
+  have "emeasure M (G i) = 0"
+    by (intro emeasure_measure_of[OF M_eq]) (auto simp add: M measure_space_def sets_extend_measure)
+  with `i\<in>I` * show ?thesis
+    by simp
+next
+  def P \<equiv> "\<lambda>\<mu>'. (\<forall>i\<in>I. \<mu>' (G i) = \<mu> i) \<and> measure_space \<Omega> (sigma_sets \<Omega> (G`I)) \<mu>'"
+  assume "\<not> (\<forall>i\<in>I. \<mu> i = 0)"
+  moreover
+  have "measure_space (space M) (sets M) \<mu>'"
+    using ms unfolding measure_space_def by auto default
+  with ms eq have "\<exists>\<mu>'. P \<mu>'"
+    unfolding P_def
+    by (intro exI[of _ \<mu>']) (auto simp add: M space_extend_measure sets_extend_measure)
+  ultimately have M_eq: "M = measure_of \<Omega> (G`I) (Eps P)"
+    by (simp add: M extend_measure_def P_def[symmetric])
+
+  from `\<exists>\<mu>'. P \<mu>'` have P: "P (Eps P)" by (rule someI_ex)
+  show "emeasure M (G i) = \<mu> i"
+  proof (subst emeasure_measure_of[OF M_eq])
+    have sets_M: "sets M = sigma_sets \<Omega> (G`I)"
+      using M_eq ms by (auto simp: sets_extend_measure)
+    then show "G i \<in> sets M" using `i \<in> I` by auto
+    show "positive (sets M) (Eps P)" "countably_additive (sets M) (Eps P)" "Eps P (G i) = \<mu> i"
+      using P `i\<in>I` by (auto simp add: sets_M measure_space_def P_def)
+  qed fact
+qed
+
+lemma emeasure_extend_measure_Pair:
+  assumes M: "M = extend_measure \<Omega> {(i, j). I i j} (\<lambda>(i, j). G i j) (\<lambda>(i, j). \<mu> i j)"
+    and eq: "\<And>i j. I i j \<Longrightarrow> \<mu>' (G i j) = \<mu> i j"
+    and ms: "\<And>i j. I i j \<Longrightarrow> G i j \<in> Pow \<Omega>" "positive (sets M) \<mu>'" "countably_additive (sets M) \<mu>'"
+    and "I i j"
+  shows "emeasure M (G i j) = \<mu> i j"
+  using emeasure_extend_measure[OF M _ _ ms(2,3), of "(i,j)"] eq ms(1) `I i j`
+  by (auto simp: subset_eq)
+
 subsection {* Sigma algebra generated by function preimages *}
 
-definition (in sigma_algebra)
-  "vimage_algebra S f = \<lparr> space = S, sets = (\<lambda>A. f -` A \<inter> S) ` sets M, \<dots> = more M \<rparr>"
-
-lemma (in sigma_algebra) in_vimage_algebra[simp]:
-  "A \<in> sets (vimage_algebra S f) \<longleftrightarrow> (\<exists>B\<in>sets M. A = f -` B \<inter> S)"
-  by (simp add: vimage_algebra_def image_iff)
+definition
+  "vimage_algebra M S f = sigma S ((\<lambda>A. f -` A \<inter> S) ` sets M)"
 
-lemma (in sigma_algebra) space_vimage_algebra[simp]:
-  "space (vimage_algebra S f) = S"
-  by (simp add: vimage_algebra_def)
-
-lemma (in sigma_algebra) sigma_algebra_preimages:
+lemma 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>")
+  assumes "f \<in> S \<rightarrow> space M"
+  shows "sigma_algebra S ((\<lambda>A. f -` A \<inter> S) ` sets M)"
+    (is "sigma_algebra _ (?F ` sets M)")
 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)"
+  fix A assume "A \<in> sets M"
+  moreover have "S - ?F A = ?F (space M - A)"
     using assms by auto
-  ultimately show "A - ?F S \<in> ?F ` sets M"
+  ultimately show "S - ?F A \<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"
+  fix A :: "nat \<Rightarrow> 'x set" assume *: "range A \<subseteq> ?F ` M"
+  have "\<forall>i. \<exists>b. b \<in> M \<and> A 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
+    have "A i \<in> ?F ` M" using * by auto
+    then show "\<exists>b. b \<in> M \<and> A 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)"
+  from choice[OF this] obtain b where b: "range b \<subseteq> M" "\<And>i. A 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
+  then have "(\<Union>i. A i) = ?F (\<Union>i. b i)" by auto
+  then show "(\<Union>i. A i) \<in> ?F ` M" using b(1) by blast
 qed
 
-lemma (in sigma_algebra) sigma_algebra_vimage:
+lemma sets_vimage_algebra[simp]:
+  "f \<in> S \<rightarrow> space M \<Longrightarrow> sets (vimage_algebra M S f) = (\<lambda>A. f -` A \<inter> S) ` sets M"
+  using sigma_algebra.sets_measure_of_eq[OF sigma_algebra_preimages, of f S M]
+  by (simp add: vimage_algebra_def)
+
+lemma space_vimage_algebra[simp]:
+  "f \<in> S \<rightarrow> space M \<Longrightarrow> space (vimage_algebra M S f) = S"
+  using sigma_algebra.space_measure_of_eq[OF sigma_algebra_preimages, of f S M]
+  by (simp add: vimage_algebra_def)
+
+lemma in_vimage_algebra[simp]:
+  "f \<in> S \<rightarrow> space M \<Longrightarrow> A \<in> sets (vimage_algebra M S f) \<longleftrightarrow> (\<exists>B\<in>sets M. A = f -` B \<inter> S)"
+  by (simp add: image_iff)
+
+lemma measurable_vimage_algebra:
   fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M"
-  shows "sigma_algebra (vimage_algebra S f)"
-proof -
-  from sigma_algebra_preimages[OF assms]
-  show ?thesis unfolding vimage_algebra_def by (auto simp: sigma_algebra_iff2)
-qed
+  shows "f \<in> measurable (vimage_algebra M S f) M"
+  unfolding measurable_def using assms by force
 
-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:
+lemma 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"
+  shows "(\<lambda>x. g (f x)) \<in> measurable (vimage_algebra M S f) M2"
 proof -
   note measurable_vimage_algebra[OF assms(2)]
   from measurable_comp[OF this assms(1)]
@@ -1033,7 +1275,7 @@
   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)
+    then show ?case by auto
   next
     case Empty then show ?case
       by (auto intro!: image_eqI[of _ _ "{}"] sigma_sets.Empty)
@@ -1060,8 +1302,6 @@
   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')
@@ -1078,23 +1318,7 @@
     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)
-  "image_space X = \<lparr> space = X`space M, sets = (\<lambda>S. X`S) ` sets M, \<dots> = more M \<rparr>"
-
-definition (in algebra)
-  "conditional_space X A = algebra.image_space (restricted_space A) X"
-
-lemma (in algebra) space_conditional_space:
-  assumes "A \<in> sets M" shows "space (conditional_space X A) = X`A"
-proof -
-  interpret r: algebra "restricted_space A" using assms by (rule restricted_algebra)
-  show ?thesis unfolding conditional_space_def r.image_space_def
-    by simp
+  qed auto
 qed
 
 subsection {* A Two-Element Series *}
@@ -1113,80 +1337,64 @@
 
 section {* Closed CDI *}
 
-definition
-  closed_cdi  where
-  "closed_cdi M \<longleftrightarrow>
-   sets M \<subseteq> Pow (space M) &
-   (\<forall>s \<in> sets M. space M - s \<in> sets M) &
-   (\<forall>A. (range A \<subseteq> sets M) & (A 0 = {}) & (\<forall>n. A n \<subseteq> A (Suc n)) \<longrightarrow>
-        (\<Union>i. A i) \<in> sets M) &
-   (\<forall>A. (range A \<subseteq> sets M) & disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
+definition closed_cdi where
+  "closed_cdi \<Omega> M \<longleftrightarrow>
+   M \<subseteq> Pow \<Omega> &
+   (\<forall>s \<in> M. \<Omega> - s \<in> M) &
+   (\<forall>A. (range A \<subseteq> M) & (A 0 = {}) & (\<forall>n. A n \<subseteq> A (Suc n)) \<longrightarrow>
+        (\<Union>i. A i) \<in> M) &
+   (\<forall>A. (range A \<subseteq> M) & disjoint_family A \<longrightarrow> (\<Union>i::nat. A i) \<in> M)"
 
 inductive_set
-  smallest_ccdi_sets :: "('a, 'b) algebra_scheme \<Rightarrow> 'a set set"
-  for M
+  smallest_ccdi_sets :: "'a set \<Rightarrow> 'a set set \<Rightarrow> 'a set set"
+  for \<Omega> M
   where
     Basic [intro]:
-      "a \<in> sets M \<Longrightarrow> a \<in> smallest_ccdi_sets M"
+      "a \<in> M \<Longrightarrow> a \<in> smallest_ccdi_sets \<Omega> M"
   | Compl [intro]:
-      "a \<in> smallest_ccdi_sets M \<Longrightarrow> space M - a \<in> smallest_ccdi_sets M"
+      "a \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> \<Omega> - a \<in> smallest_ccdi_sets \<Omega> M"
   | Inc:
-      "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> A 0 = {} \<Longrightarrow> (\<And>n. A n \<subseteq> A (Suc n))
-       \<Longrightarrow> (\<Union>i. A i) \<in> smallest_ccdi_sets M"
+      "range A \<in> Pow(smallest_ccdi_sets \<Omega> M) \<Longrightarrow> A 0 = {} \<Longrightarrow> (\<And>n. A n \<subseteq> A (Suc n))
+       \<Longrightarrow> (\<Union>i. A i) \<in> smallest_ccdi_sets \<Omega> M"
   | Disj:
-      "range A \<in> Pow(smallest_ccdi_sets M) \<Longrightarrow> disjoint_family A
-       \<Longrightarrow> (\<Union>i::nat. A i) \<in> smallest_ccdi_sets M"
-
+      "range A \<in> Pow(smallest_ccdi_sets \<Omega> M) \<Longrightarrow> disjoint_family A
+       \<Longrightarrow> (\<Union>i::nat. A i) \<in> smallest_ccdi_sets \<Omega> M"
 
-definition
-  smallest_closed_cdi  where
-  "smallest_closed_cdi M = (|space = space M, sets = smallest_ccdi_sets M|)"
+lemma (in subset_class) smallest_closed_cdi1: "M \<subseteq> smallest_ccdi_sets \<Omega> M"
+  by auto
 
-lemma space_smallest_closed_cdi [simp]:
-     "space (smallest_closed_cdi M) = space M"
-  by (simp add: smallest_closed_cdi_def)
-
-lemma (in algebra) smallest_closed_cdi1: "sets M \<subseteq> sets (smallest_closed_cdi M)"
-  by (auto simp add: smallest_closed_cdi_def)
-
-lemma (in algebra) smallest_ccdi_sets:
-     "smallest_ccdi_sets M \<subseteq> Pow (space M)"
+lemma (in subset_class) smallest_ccdi_sets: "smallest_ccdi_sets \<Omega> M \<subseteq> Pow \<Omega>"
   apply (rule subsetI)
   apply (erule smallest_ccdi_sets.induct)
   apply (auto intro: range_subsetD dest: sets_into_space)
   done
 
-lemma (in algebra) smallest_closed_cdi2: "closed_cdi (smallest_closed_cdi M)"
-  apply (auto simp add: closed_cdi_def smallest_closed_cdi_def smallest_ccdi_sets)
+lemma (in subset_class) smallest_closed_cdi2: "closed_cdi \<Omega> (smallest_ccdi_sets \<Omega> M)"
+  apply (auto simp add: closed_cdi_def smallest_ccdi_sets)
   apply (blast intro: smallest_ccdi_sets.Inc smallest_ccdi_sets.Disj) +
   done
 
-lemma (in algebra) smallest_closed_cdi3:
-     "sets (smallest_closed_cdi M) \<subseteq> Pow (space M)"
-  by (simp add: smallest_closed_cdi_def smallest_ccdi_sets)
-
-lemma closed_cdi_subset: "closed_cdi M \<Longrightarrow> sets M \<subseteq> Pow (space M)"
+lemma closed_cdi_subset: "closed_cdi \<Omega> M \<Longrightarrow> M \<subseteq> Pow \<Omega>"
   by (simp add: closed_cdi_def)
 
-lemma closed_cdi_Compl: "closed_cdi M \<Longrightarrow> s \<in> sets M \<Longrightarrow> space M - s \<in> sets M"
+lemma closed_cdi_Compl: "closed_cdi \<Omega> M \<Longrightarrow> s \<in> M \<Longrightarrow> \<Omega> - s \<in> M"
   by (simp add: closed_cdi_def)
 
 lemma closed_cdi_Inc:
-     "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) \<Longrightarrow>
-        (\<Union>i. A i) \<in> sets M"
+  "closed_cdi \<Omega> M \<Longrightarrow> range A \<subseteq> M \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n)) \<Longrightarrow> (\<Union>i. A i) \<in> M"
   by (simp add: closed_cdi_def)
 
 lemma closed_cdi_Disj:
-     "closed_cdi M \<Longrightarrow> range A \<subseteq> sets M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
+  "closed_cdi \<Omega> M \<Longrightarrow> range A \<subseteq> M \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
   by (simp add: closed_cdi_def)
 
 lemma closed_cdi_Un:
-  assumes cdi: "closed_cdi M" and empty: "{} \<in> sets M"
-      and A: "A \<in> sets M" and B: "B \<in> sets M"
+  assumes cdi: "closed_cdi \<Omega> M" and empty: "{} \<in> M"
+      and A: "A \<in> M" and B: "B \<in> M"
       and disj: "A \<inter> B = {}"
-    shows "A \<union> B \<in> sets M"
+    shows "A \<union> B \<in> M"
 proof -
-  have ra: "range (binaryset A B) \<subseteq> sets M"
+  have ra: "range (binaryset A B) \<subseteq> M"
    by (simp add: range_binaryset_eq empty A B)
  have di:  "disjoint_family (binaryset A B)" using disj
    by (simp add: disjoint_family_on_def binaryset_def Int_commute)
@@ -1196,11 +1404,11 @@
 qed
 
 lemma (in algebra) smallest_ccdi_sets_Un:
-  assumes A: "A \<in> smallest_ccdi_sets M" and B: "B \<in> smallest_ccdi_sets M"
+  assumes A: "A \<in> smallest_ccdi_sets \<Omega> M" and B: "B \<in> smallest_ccdi_sets \<Omega> M"
       and disj: "A \<inter> B = {}"
-    shows "A \<union> B \<in> smallest_ccdi_sets M"
+    shows "A \<union> B \<in> smallest_ccdi_sets \<Omega> M"
 proof -
-  have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets M)"
+  have ra: "range (binaryset A B) \<in> Pow (smallest_ccdi_sets \<Omega> M)"
     by (simp add: range_binaryset_eq  A B smallest_ccdi_sets.Basic)
   have di:  "disjoint_family (binaryset A B)" using disj
     by (simp add: disjoint_family_on_def binaryset_def Int_commute)
@@ -1210,33 +1418,32 @@
 qed
 
 lemma (in algebra) smallest_ccdi_sets_Int1:
-  assumes a: "a \<in> sets M"
-  shows "b \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M"
+  assumes a: "a \<in> M"
+  shows "b \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets \<Omega> M"
 proof (induct rule: smallest_ccdi_sets.induct)
   case (Basic x)
   thus ?case
     by (metis a Int smallest_ccdi_sets.Basic)
 next
   case (Compl x)
-  have "a \<inter> (space M - x) = space M - ((space M - a) \<union> (a \<inter> x))"
+  have "a \<inter> (\<Omega> - x) = \<Omega> - ((\<Omega> - a) \<union> (a \<inter> x))"
     by blast
-  also have "... \<in> smallest_ccdi_sets M"
+  also have "... \<in> smallest_ccdi_sets \<Omega> M"
     by (metis smallest_ccdi_sets.Compl a Compl(2) Diff_Int2 Diff_Int_distrib2
-           Diff_disjoint Int_Diff Int_empty_right Un_commute
-           smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl
-           smallest_ccdi_sets_Un)
+           Diff_disjoint Int_Diff Int_empty_right smallest_ccdi_sets_Un
+           smallest_ccdi_sets.Basic smallest_ccdi_sets.Compl)
   finally show ?case .
 next
   case (Inc A)
   have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
     by blast
-  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Inc
+  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Inc
     by blast
   moreover have "(\<lambda>i. a \<inter> A i) 0 = {}"
     by (simp add: Inc)
   moreover have "!!n. (\<lambda>i. a \<inter> A i) n \<subseteq> (\<lambda>i. a \<inter> A i) (Suc n)" using Inc
     by blast
-  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M"
+  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets \<Omega> M"
     by (rule smallest_ccdi_sets.Inc)
   show ?case
     by (metis 1 2)
@@ -1244,11 +1451,11 @@
   case (Disj A)
   have 1: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) = a \<inter> (\<Union>i. A i)"
     by blast
-  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets M)" using Disj
+  have "range (\<lambda>i. a \<inter> A i) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Disj
     by blast
   moreover have "disjoint_family (\<lambda>i. a \<inter> A i)" using Disj
     by (auto simp add: disjoint_family_on_def)
-  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets M"
+  ultimately have 2: "(\<Union>i. (\<lambda>i. a \<inter> A i) i) \<in> smallest_ccdi_sets \<Omega> M"
     by (rule smallest_ccdi_sets.Disj)
   show ?case
     by (metis 1 2)
@@ -1256,17 +1463,17 @@
 
 
 lemma (in algebra) smallest_ccdi_sets_Int:
-  assumes b: "b \<in> smallest_ccdi_sets M"
-  shows "a \<in> smallest_ccdi_sets M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets M"
+  assumes b: "b \<in> smallest_ccdi_sets \<Omega> M"
+  shows "a \<in> smallest_ccdi_sets \<Omega> M \<Longrightarrow> a \<inter> b \<in> smallest_ccdi_sets \<Omega> M"
 proof (induct rule: smallest_ccdi_sets.induct)
   case (Basic x)
   thus ?case
     by (metis b smallest_ccdi_sets_Int1)
 next
   case (Compl x)
-  have "(space M - x) \<inter> b = space M - (x \<inter> b \<union> (space M - b))"
+  have "(\<Omega> - x) \<inter> b = \<Omega> - (x \<inter> b \<union> (\<Omega> - b))"
     by blast
-  also have "... \<in> smallest_ccdi_sets M"
+  also have "... \<in> smallest_ccdi_sets \<Omega> M"
     by (metis Compl(2) Diff_disjoint Int_Diff Int_commute Int_empty_right b
            smallest_ccdi_sets.Compl smallest_ccdi_sets_Un)
   finally show ?case .
@@ -1274,13 +1481,13 @@
   case (Inc A)
   have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
     by blast
-  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Inc
+  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Inc
     by blast
   moreover have "(\<lambda>i. A i \<inter> b) 0 = {}"
     by (simp add: Inc)
   moreover have "!!n. (\<lambda>i. A i \<inter> b) n \<subseteq> (\<lambda>i. A i \<inter> b) (Suc n)" using Inc
     by blast
-  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M"
+  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets \<Omega> M"
     by (rule smallest_ccdi_sets.Inc)
   show ?case
     by (metis 1 2)
@@ -1288,39 +1495,34 @@
   case (Disj A)
   have 1: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) = (\<Union>i. A i) \<inter> b"
     by blast
-  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets M)" using Disj
+  have "range (\<lambda>i. A i \<inter> b) \<in> Pow(smallest_ccdi_sets \<Omega> M)" using Disj
     by blast
   moreover have "disjoint_family (\<lambda>i. A i \<inter> b)" using Disj
     by (auto simp add: disjoint_family_on_def)
-  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets M"
+  ultimately have 2: "(\<Union>i. (\<lambda>i. A i \<inter> b) i) \<in> smallest_ccdi_sets \<Omega> M"
     by (rule smallest_ccdi_sets.Disj)
   show ?case
     by (metis 1 2)
 qed
 
-lemma (in algebra) sets_smallest_closed_cdi_Int:
-   "a \<in> sets (smallest_closed_cdi M) \<Longrightarrow> b \<in> sets (smallest_closed_cdi M)
-    \<Longrightarrow> a \<inter> b \<in> sets (smallest_closed_cdi M)"
-  by (simp add: smallest_ccdi_sets_Int smallest_closed_cdi_def)
-
 lemma (in algebra) sigma_property_disjoint_lemma:
-  assumes sbC: "sets M \<subseteq> C"
-      and ccdi: "closed_cdi (|space = space M, sets = C|)"
-  shows "sigma_sets (space M) (sets M) \<subseteq> C"
+  assumes sbC: "M \<subseteq> C"
+      and ccdi: "closed_cdi \<Omega> C"
+  shows "sigma_sets \<Omega> M \<subseteq> C"
 proof -
-  have "smallest_ccdi_sets M \<in> {B . sets M \<subseteq> B \<and> sigma_algebra (|space = space M, sets = B|)}"
+  have "smallest_ccdi_sets \<Omega> M \<in> {B . M \<subseteq> B \<and> sigma_algebra \<Omega> B}"
     apply (auto simp add: sigma_algebra_disjoint_iff algebra_iff_Int
             smallest_ccdi_sets_Int)
     apply (metis Union_Pow_eq Union_upper subsetD smallest_ccdi_sets)
     apply (blast intro: smallest_ccdi_sets.Disj)
     done
-  hence "sigma_sets (space M) (sets M) \<subseteq> smallest_ccdi_sets M"
+  hence "sigma_sets (\<Omega>) (M) \<subseteq> smallest_ccdi_sets \<Omega> M"
     by clarsimp
-       (drule sigma_algebra.sigma_sets_subset [where a="sets M"], auto)
+       (drule sigma_algebra.sigma_sets_subset [where a="M"], auto)
   also have "...  \<subseteq> C"
     proof
       fix x
-      assume x: "x \<in> smallest_ccdi_sets M"
+      assume x: "x \<in> smallest_ccdi_sets \<Omega> M"
       thus "x \<in> C"
         proof (induct rule: smallest_ccdi_sets.induct)
           case (Basic x)
@@ -1344,21 +1546,21 @@
 qed
 
 lemma (in algebra) sigma_property_disjoint:
-  assumes sbC: "sets M \<subseteq> C"
-      and compl: "!!s. s \<in> C \<inter> sigma_sets (space M) (sets M) \<Longrightarrow> space M - s \<in> C"
-      and inc: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M)
+  assumes sbC: "M \<subseteq> C"
+      and compl: "!!s. s \<in> C \<inter> sigma_sets (\<Omega>) (M) \<Longrightarrow> \<Omega> - s \<in> C"
+      and inc: "!!A. range A \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)
                      \<Longrightarrow> A 0 = {} \<Longrightarrow> (!!n. A n \<subseteq> A (Suc n))
                      \<Longrightarrow> (\<Union>i. A i) \<in> C"
-      and disj: "!!A. range A \<subseteq> C \<inter> sigma_sets (space M) (sets M)
+      and disj: "!!A. range A \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)
                       \<Longrightarrow> disjoint_family A \<Longrightarrow> (\<Union>i::nat. A i) \<in> C"
-  shows "sigma_sets (space M) (sets M) \<subseteq> C"
+  shows "sigma_sets (\<Omega>) (M) \<subseteq> C"
 proof -
-  have "sigma_sets (space M) (sets M) \<subseteq> C \<inter> sigma_sets (space M) (sets M)"
+  have "sigma_sets (\<Omega>) (M) \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)"
     proof (rule sigma_property_disjoint_lemma)
-      show "sets M \<subseteq> C \<inter> sigma_sets (space M) (sets M)"
+      show "M \<subseteq> C \<inter> sigma_sets (\<Omega>) (M)"
         by (metis Int_greatest Set.subsetI sbC sigma_sets.Basic)
     next
-      show "closed_cdi \<lparr>space = space M, sets = C \<inter> sigma_sets (space M) (sets M)\<rparr>"
+      show "closed_cdi \<Omega> (C \<inter> sigma_sets (\<Omega>) (M))"
         by (simp add: closed_cdi_def compl inc disj)
            (metis PowI Set.subsetI le_infI2 sigma_sets_into_sp space_closed
              IntE sigma_sets.Compl range_subsetD sigma_sets.Union)
@@ -1370,97 +1572,97 @@
 section {* Dynkin systems *}
 
 locale dynkin_system = subset_class +
-  assumes 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"
+  assumes space: "\<Omega> \<in> M"
+    and   compl[intro!]: "\<And>A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
+    and   UN[intro!]: "\<And>A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
+                           \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
 
-lemma (in dynkin_system) empty[intro, simp]: "{} \<in> sets M"
-  using space compl[of "space M"] by simp
+lemma (in dynkin_system) empty[intro, simp]: "{} \<in> M"
+  using space compl[of "\<Omega>"] by simp
 
 lemma (in dynkin_system) diff:
-  assumes sets: "D \<in> sets M" "E \<in> sets M" and "D \<subseteq> E"
-  shows "E - D \<in> sets M"
+  assumes sets: "D \<in> M" "E \<in> M" and "D \<subseteq> E"
+  shows "E - D \<in> M"
 proof -
-  let ?f = "\<lambda>x. if x = 0 then D else if x = Suc 0 then space M - E else {}"
-  have "range ?f = {D, space M - E, {}}"
+  let ?f = "\<lambda>x. if x = 0 then D else if x = Suc 0 then \<Omega> - E else {}"
+  have "range ?f = {D, \<Omega> - E, {}}"
     by (auto simp: image_iff)
-  moreover have "D \<union> (space M - E) = (\<Union>i. ?f i)"
+  moreover have "D \<union> (\<Omega> - 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 `D \<in> M`[THEN sets_into_space] `D \<subseteq> E` by auto
+  ultimately have "\<Omega> - (D \<union> (\<Omega> - E)) \<in> M"
     using sets by auto
-  also have "space M - (D \<union> (space M - E)) = E - D"
+  also have "\<Omega> - (D \<union> (\<Omega> - E)) = E - D"
     using assms sets_into_space by auto
   finally show ?thesis .
 qed
 
 lemma dynkin_systemI:
-  assumes "\<And> A. A \<in> 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"
+  assumes "\<And> A. A \<in> M \<Longrightarrow> A \<subseteq> \<Omega>" "\<Omega> \<in> M"
+  assumes "\<And> A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
+  assumes "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
+          \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
+  shows "dynkin_system \<Omega> M"
   using assms by (auto simp: dynkin_system_def dynkin_system_axioms_def subset_class_def)
 
 lemma dynkin_systemI':
-  assumes 1: "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M"
-  assumes empty: "{} \<in> sets M"
-  assumes Diff: "\<And> A. A \<in> sets M \<Longrightarrow> space M - A \<in> sets M"
-  assumes 2: "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> sets M
-          \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
-  shows "dynkin_system M"
+  assumes 1: "\<And> A. A \<in> M \<Longrightarrow> A \<subseteq> \<Omega>"
+  assumes empty: "{} \<in> M"
+  assumes Diff: "\<And> A. A \<in> M \<Longrightarrow> \<Omega> - A \<in> M"
+  assumes 2: "\<And> A. disjoint_family A \<Longrightarrow> range A \<subseteq> M
+          \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
+  shows "dynkin_system \<Omega> M"
 proof -
-  from Diff[OF empty] have "space M \<in> sets M" by auto
+  from Diff[OF empty] have "\<Omega> \<in> M" by auto
   from 1 this Diff 2 show ?thesis
     by (intro dynkin_systemI) auto
 qed
 
 lemma dynkin_system_trivial:
-  shows "dynkin_system \<lparr> space = A, sets = Pow A \<rparr>"
+  shows "dynkin_system A (Pow A)"
   by (rule dynkin_systemI) auto
 
 lemma sigma_algebra_imp_dynkin_system:
-  assumes "sigma_algebra M" shows "dynkin_system M"
+  assumes "sigma_algebra \<Omega> M" shows "dynkin_system \<Omega> M"
 proof -
-  interpret sigma_algebra M by fact
+  interpret sigma_algebra \<Omega> M by fact
   show ?thesis using sets_into_space by (fastforce 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)"
+definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)"
 
 lemma (in algebra) Int_stable: "Int_stable M"
   unfolding Int_stable_def by auto
 
 lemma Int_stableI:
-  "(\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<inter> b \<in> A) \<Longrightarrow> Int_stable \<lparr> space = \<Omega>, sets = A \<rparr>"
+  "(\<And>a b. a \<in> A \<Longrightarrow> b \<in> A \<Longrightarrow> a \<inter> b \<in> A) \<Longrightarrow> Int_stable A"
   unfolding Int_stable_def by auto
 
 lemma Int_stableD:
-  "Int_stable M \<Longrightarrow> a \<in> sets M \<Longrightarrow> b \<in> sets M \<Longrightarrow> a \<inter> b \<in> sets M"
+  "Int_stable M \<Longrightarrow> a \<in> M \<Longrightarrow> b \<in> M \<Longrightarrow> a \<inter> b \<in> M"
   unfolding Int_stable_def by auto
 
 lemma (in dynkin_system) sigma_algebra_eq_Int_stable:
-  "sigma_algebra M \<longleftrightarrow> Int_stable M"
+  "sigma_algebra \<Omega> M \<longleftrightarrow> Int_stable M"
 proof
-  assume "sigma_algebra M" then show "Int_stable M"
+  assume "sigma_algebra \<Omega> M" then show "Int_stable M"
     unfolding sigma_algebra_def using algebra.Int_stable by auto
 next
   assume "Int_stable M"
-  show "sigma_algebra M"
+  show "sigma_algebra \<Omega> M"
     unfolding sigma_algebra_disjoint_iff algebra_iff_Un
   proof (intro conjI ballI allI impI)
-    show "sets M \<subseteq> Pow (space M)" using sets_into_space by auto
+    show "M \<subseteq> Pow (\<Omega>)" 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"
+    fix A B assume "A \<in> M" "B \<in> M"
+    then have "A \<union> B = \<Omega> - ((\<Omega> - A) \<inter> (\<Omega> - B))"
+              "\<Omega> - A \<in> M" "\<Omega> - B \<in> M"
       using sets_into_space by auto
-    then show "A \<union> B \<in> sets M"
+    then show "A \<union> B \<in> M"
       using `Int_stable M` unfolding Int_stable_def by auto
   qed auto
 qed
@@ -1468,217 +1670,148 @@
 subsection "Smallest Dynkin systems"
 
 definition dynkin where
-  "dynkin M = \<lparr> space = space M,
-     sets = \<Inter>{D. dynkin_system \<lparr> space = space M, sets = D \<rparr> \<and> sets M \<subseteq> D},
-     \<dots> = more M \<rparr>"
+  "dynkin \<Omega> M =  (\<Inter>{D. dynkin_system \<Omega> D \<and> M \<subseteq> D})"
 
 lemma dynkin_system_dynkin:
-  assumes "sets M \<subseteq> Pow (space M)"
-  shows "dynkin_system (dynkin M)"
+  assumes "M \<subseteq> Pow (\<Omega>)"
+  shows "dynkin_system \<Omega> (dynkin \<Omega> M)"
 proof (rule dynkin_systemI)
-  fix A assume "A \<in> sets (dynkin M)"
+  fix A assume "A \<in> dynkin \<Omega> M"
   moreover
-  { fix D assume "A \<in> D" and d: "dynkin_system \<lparr> space = space M, sets = D \<rparr>"
-    then have "A \<subseteq> space M" by (auto simp: dynkin_system_def subset_class_def) }
-  moreover have "{D. dynkin_system \<lparr> space = space M, sets = D\<rparr> \<and> sets M \<subseteq> D} \<noteq> {}"
+  { fix D assume "A \<in> D" and d: "dynkin_system \<Omega> D"
+    then have "A \<subseteq> \<Omega>" by (auto simp: dynkin_system_def subset_class_def) }
+  moreover have "{D. dynkin_system \<Omega> D \<and> M \<subseteq> D} \<noteq> {}"
     using assms dynkin_system_trivial by fastforce
-  ultimately show "A \<subseteq> space (dynkin M)"
+  ultimately show "A \<subseteq> \<Omega>"
     unfolding dynkin_def using assms
-    by simp (metis dynkin_system_def subset_class_def in_mono)
+    by auto
 next
-  show "space (dynkin M) \<in> sets (dynkin M)"
+  show "\<Omega> \<in> dynkin \<Omega> M"
     unfolding dynkin_def using dynkin_system.space by fastforce
 next
-  fix A assume "A \<in> sets (dynkin M)"
-  then show "space (dynkin M) - A \<in> sets (dynkin M)"
+  fix A assume "A \<in> dynkin \<Omega> M"
+  then show "\<Omega> - A \<in> dynkin \<Omega> M"
     unfolding dynkin_def using dynkin_system.compl by force
 next
   fix A :: "nat \<Rightarrow> 'a set"
-  assume A: "disjoint_family A" "range A \<subseteq> sets (dynkin M)"
-  show "(\<Union>i. A i) \<in> sets (dynkin M)" unfolding dynkin_def
+  assume A: "disjoint_family A" "range A \<subseteq> dynkin \<Omega> M"
+  show "(\<Union>i. A i) \<in> dynkin \<Omega> M" unfolding dynkin_def
   proof (simp, safe)
-    fix D assume "dynkin_system \<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>"
+    fix D assume "dynkin_system \<Omega> D" "M \<subseteq> D"
+    with A have "(\<Union>i. A i) \<in> D"
       by (intro dynkin_system.UN) (auto simp: dynkin_def)
     then show "(\<Union>i. A i) \<in> D" by auto
   qed
 qed
 
-lemma dynkin_Basic[intro]:
-  "A \<in> sets M \<Longrightarrow> A \<in> sets (dynkin M)"
-  unfolding dynkin_def by auto
-
-lemma dynkin_space[simp]:
-  "space (dynkin M) = space M"
+lemma dynkin_Basic[intro]: "A \<in> M \<Longrightarrow> A \<in> dynkin \<Omega> M"
   unfolding dynkin_def by auto
 
 lemma (in dynkin_system) restricted_dynkin_system:
-  assumes "D \<in> sets M"
-  shows "dynkin_system \<lparr> space = space M,
-                         sets = {Q. Q \<subseteq> space M \<and> Q \<inter> D \<in> sets M} \<rparr>"
+  assumes "D \<in> M"
+  shows "dynkin_system \<Omega> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
 proof (rule dynkin_systemI, simp_all)
-  have "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
+  have "\<Omega> \<inter> D = D"
+    using `D \<in> M` sets_into_space by auto
+  then show "\<Omega> \<inter> D \<in> M"
+    using `D \<in> 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)"
+  fix A assume "A \<subseteq> \<Omega> \<and> A \<inter> D \<in> M"
+  moreover have "(\<Omega> - A) \<inter> D = (\<Omega> - (A \<inter> D)) - (\<Omega> - D)"
     by auto
-  ultimately show "space M - A \<subseteq> space M \<and> (space M - A) \<inter> D \<in> sets M"
-    using  `D \<in> sets M` by (auto intro: diff)
+  ultimately show "\<Omega> - A \<subseteq> \<Omega> \<and> (\<Omega> - A) \<inter> D \<in> M"
+    using  `D \<in> 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)"
+  assume "disjoint_family A" "range A \<subseteq> {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> D \<in> M}"
+  then have "\<And>i. A i \<subseteq> \<Omega>" "disjoint_family (\<lambda>i. A i \<inter> D)"
+    "range (\<lambda>i. A i \<inter> D) \<subseteq> M" "(\<Union>x. A x) \<inter> D = (\<Union>x. A x \<inter> D)"
     by ((fastforce simp: disjoint_family_on_def)+)
-  then show "(\<Union>x. A x) \<subseteq> space M \<and> (\<Union>x. A x) \<inter> D \<in> sets M"
+  then show "(\<Union>x. A x) \<subseteq> \<Omega> \<and> (\<Union>x. A x) \<inter> D \<in> M"
     by (auto simp del: UN_simps)
 qed
 
 lemma (in dynkin_system) dynkin_subset:
-  assumes "sets N \<subseteq> sets M"
-  assumes "space N = space M"
-  shows "sets (dynkin N) \<subseteq> sets M"
+  assumes "N \<subseteq> M"
+  shows "dynkin \<Omega> N \<subseteq> M"
 proof -
-  have "dynkin_system M" by default
-  then have "dynkin_system \<lparr>space = space N, sets = sets M \<rparr>"
+  have "dynkin_system \<Omega> M" by default
+  then have "dynkin_system \<Omega> M"
     using assms unfolding dynkin_system_def dynkin_system_axioms_def subset_class_def by simp
-  with `sets N \<subseteq> sets M` show ?thesis by (auto simp add: dynkin_def)
+  with `N \<subseteq> M` show ?thesis by (auto simp add: dynkin_def)
 qed
 
 lemma sigma_eq_dynkin:
-  assumes sets: "sets M \<subseteq> Pow (space M)"
+  assumes sets: "M \<subseteq> Pow \<Omega>"
   assumes "Int_stable M"
-  shows "sigma M = dynkin M"
+  shows "sigma_sets \<Omega> M = dynkin \<Omega> M"
 proof -
-  have "sets (dynkin M) \<subseteq> sigma_sets (space M) (sets M)"
+  have "dynkin \<Omega> M \<subseteq> sigma_sets (\<Omega>) (M)"
     using sigma_algebra_imp_dynkin_system
-    unfolding dynkin_def sigma_def sigma_sets_least_sigma_algebra[OF sets] by auto
+    unfolding dynkin_def sigma_sets_least_sigma_algebra[OF sets] by auto
   moreover
-  interpret dynkin_system "dynkin M"
+  interpret dynkin_system \<Omega> "dynkin \<Omega> M"
     using dynkin_system_dynkin[OF sets] .
-  have "sigma_algebra (dynkin M)"
+  have "sigma_algebra \<Omega> (dynkin \<Omega> M)"
     unfolding sigma_algebra_eq_Int_stable Int_stable_def
   proof (intro ballI)
-    fix A B assume "A \<in> sets (dynkin M)" "B \<in> sets (dynkin M)"
-    let ?D = "\<lambda>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)"
+    fix A B assume "A \<in> dynkin \<Omega> M" "B \<in> dynkin \<Omega> M"
+    let ?D = "\<lambda>E. {Q. Q \<subseteq> \<Omega> \<and> Q \<inter> E \<in> dynkin \<Omega> M}"
+    have "M \<subseteq> ?D B"
     proof
-      fix E assume "E \<in> sets M"
-      then have "sets M \<subseteq> sets (?D E)" "E \<in> sets (dynkin M)"
+      fix E assume "E \<in> M"
+      then have "M \<subseteq> ?D E" "E \<in> dynkin \<Omega> 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)`
+      then have "dynkin \<Omega> M \<subseteq> ?D E"
+        using restricted_dynkin_system `E \<in> dynkin \<Omega> 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)"
+      then have "B \<in> ?D E"
+        using `B \<in> dynkin \<Omega> M` by auto
+      then have "E \<inter> B \<in> dynkin \<Omega> M"
         by (subst Int_commute) simp
-      then show "E \<in> sets (?D B)"
-        using sets `E \<in> sets M` by auto
+      then show "E \<in> ?D B"
+        using sets `E \<in> M` by auto
     qed
-    then have "sets (dynkin M) \<subseteq> sets (?D B)"
-      using restricted_dynkin_system `B \<in> sets (dynkin M)`
+    then have "dynkin \<Omega> M \<subseteq> ?D B"
+      using restricted_dynkin_system `B \<in> dynkin \<Omega> 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
+    then show "A \<inter> B \<in> dynkin \<Omega> M"
+      using `A \<in> dynkin \<Omega> 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
+  from sigma_algebra.sigma_sets_subset[OF this, of "M"]
+  have "sigma_sets (\<Omega>) (M) \<subseteq> dynkin \<Omega> M" by auto
+  ultimately have "sigma_sets (\<Omega>) (M) = dynkin \<Omega> M" by auto
   then show ?thesis
-    by (auto intro!: algebra.equality simp: sigma_def dynkin_def)
+    by (auto simp: dynkin_def)
 qed
 
 lemma (in dynkin_system) dynkin_idem:
-  "dynkin M = M"
+  "dynkin \<Omega> M = M"
 proof -
-  have "sets (dynkin M) = sets M"
+  have "dynkin \<Omega> M = M"
   proof
-    show "sets M \<subseteq> sets (dynkin M)"
+    show "M \<subseteq> dynkin \<Omega> M"
       using dynkin_Basic by auto
-    show "sets (dynkin M) \<subseteq> sets M"
+    show "dynkin \<Omega> M \<subseteq> M"
       by (intro dynkin_subset) auto
   qed
   then show ?thesis
-    by (auto intro!: algebra.equality simp: dynkin_def)
+    by (auto simp: dynkin_def)
 qed
 
 lemma (in dynkin_system) dynkin_lemma:
   assumes "Int_stable E"
-  and E: "sets E \<subseteq> sets M" "space E = space M" "sets M \<subseteq> sets (sigma E)"
-  shows "sets (sigma E) = sets M"
+  and E: "E \<subseteq> M" "M \<subseteq> sigma_sets \<Omega> E"
+  shows "sigma_sets \<Omega> E = M"
 proof -
-  have "sets E \<subseteq> Pow (space E)"
+  have "E \<subseteq> Pow \<Omega>"
     using E sets_into_space by force
-  then have "sigma E = dynkin E"
+  then have "sigma_sets \<Omega> E = dynkin \<Omega> E"
     using `Int_stable E` by (rule sigma_eq_dynkin)
-  moreover then have "sets (dynkin E) = sets M"
-    using assms dynkin_subset[OF E(1,2)] by simp
+  moreover then have "dynkin \<Omega> E = M"
+    using assms dynkin_subset[OF E(1)] by simp
   ultimately show ?thesis
-    using assms by (auto intro!: algebra.equality simp: dynkin_def)
-qed
-
-subsection "Sigma algebras on finite sets"
-
-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 image_space_def)
-  then show "S \<in> sets (image_space X)"
-    by (auto simp: image_space_def)
-qed
-
-lemma measurable_sigma_sigma:
-  assumes M: "sets M \<subseteq> Pow (space M)" and N: "sets N \<subseteq> Pow (space N)"
-  shows "f \<in> measurable M N \<Longrightarrow> f \<in> measurable (sigma M) (sigma N)"
-  using sigma_algebra.measurable_subset[OF sigma_algebra_sigma[OF M], of N]
-  using measurable_up_sigma[of M N] N by auto
-
-lemma (in sigma_algebra) measurable_Least:
-  assumes meas: "\<And>i::nat. {x\<in>space M. P i x} \<in> sets M"
-  shows "(\<lambda>x. LEAST j. P j x) -` A \<inter> space M \<in> sets M"
-proof -
-  { fix i have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M \<in> sets M"
-    proof cases
-      assume i: "(LEAST j. False) = i"
-      have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
-        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x})) \<union> (space M - (\<Union>i. {x\<in>space M. P i x}))"
-        by (simp add: set_eq_iff, safe)
-           (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality)
-      with meas show ?thesis
-        by (auto intro!: Int)
-    next
-      assume i: "(LEAST j. False) \<noteq> i"
-      then have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
-        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x}))"
-      proof (simp add: set_eq_iff, safe)
-        fix x assume neq: "(LEAST j. False) \<noteq> (LEAST j. P j x)"
-        have "\<exists>j. P j x"
-          by (rule ccontr) (insert neq, auto)
-        then show "P (LEAST j. P j x) x" by (rule LeastI_ex)
-      qed (auto dest: Least_le intro!: Least_equality)
-      with meas show ?thesis
-        by (auto intro!: Int)
-    qed }
-  then have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) \<in> sets M"
-    by (intro countable_UN) auto
-  moreover have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) =
-    (\<lambda>x. LEAST j. P j x) -` A \<inter> space M" by auto
-  ultimately show ?thesis by auto
+    using assms by (auto simp: dynkin_def)
 qed
 
 end
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -8,29 +8,6 @@
 lemma Ex1_eq: "\<exists>! x. P x \<Longrightarrow> P x \<Longrightarrow> P y \<Longrightarrow> x = y"
   by auto
 
-locale finite_space =
-  fixes S :: "'a set"
-  assumes finite[simp]: "finite S"
-  and not_empty[simp]: "S \<noteq> {}"
-
-definition (in finite_space) "M = \<lparr> space = S, sets = Pow S,
-  measure = \<lambda>A. ereal (card A / card S) \<rparr>"
-
-lemma (in finite_space)
-  shows space_M[simp]: "space M = S"
-  and sets_M[simp]: "sets M = Pow S"
-  by (simp_all add: M_def)
-
-sublocale finite_space \<subseteq> finite_measure_space M
-  by (rule finite_measure_spaceI)
-     (simp_all add: M_def real_of_nat_def)
-
-sublocale finite_space \<subseteq> information_space M 2
-  by default (simp_all add: M_def one_ereal_def)
-
-lemma (in finite_space) \<mu>'_eq[simp]: "\<mu>' A = (if A \<subseteq> S then card A / card S else 0)"
-  unfolding \<mu>'_def by (auto simp: M_def)
-
 section "Define the state space"
 
 text {*
@@ -68,7 +45,7 @@
 definition "dining_cryptographers =
   ({None} \<union> Some ` {0..<n}) \<times> {xs :: bool list. length xs = n}"
 definition "payer dc = fst dc"
-definition coin :: "(nat option \<times> bool list) => nat \<Rightarrow> bool" where
+definition coin :: "(nat option \<times> bool list) \<Rightarrow> nat \<Rightarrow> bool" where
   "coin dc c = snd dc ! (c mod n)"
 definition "inversion dc =
   map (\<lambda>c. (payer dc = Some c) \<noteq> (coin dc c \<noteq> coin dc (c + 1))) [0..<n]"
@@ -417,15 +394,12 @@
 
 end
 
+sublocale dining_cryptographers_space \<subseteq> prob_space "uniform_count_measure dc_crypto"
+  by (rule prob_space_uniform_count_measure[OF finite_dc_crypto])
+     (insert n_gt_3, auto simp: dc_crypto intro: exI[of _ "replicate n True"])
 
-sublocale
-  dining_cryptographers_space \<subseteq> finite_space "dc_crypto"
-proof
-  show "finite dc_crypto" using finite_dc_crypto .
-  show "dc_crypto \<noteq> {}"
-    unfolding dc_crypto
-    using n_gt_3 by (auto intro: exI[of _ "replicate n True"])
-qed
+sublocale dining_cryptographers_space \<subseteq> information_space "uniform_count_measure dc_crypto" 2
+  by default auto
 
 notation (in dining_cryptographers_space)
   mutual_information_Pow ("\<I>'( _ ; _ ')")
@@ -438,71 +412,64 @@
 
 theorem (in dining_cryptographers_space)
   "\<I>( inversion ; payer ) = 0"
-proof -
+proof (rule mutual_information_eq_0_simple)
   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
 
-  let ?dIP = "distribution (\<lambda>x. (inversion x, payer x))"
-  let ?dP = "distribution payer"
-  let ?dI = "distribution inversion"
+  show inversion: "simple_distributed (uniform_count_measure dc_crypto) inversion (\<lambda>x. 2 / 2^n)"
+  proof (rule simple_distributedI)
+    show "simple_function (uniform_count_measure dc_crypto) inversion"
+      using finite_dc_crypto
+      by (auto simp: simple_function_def space_uniform_count_measure sets_uniform_count_measure)
+    fix x assume "x \<in> inversion ` space (uniform_count_measure dc_crypto)"
+    moreover have "inversion -` {x} \<inter> dc_crypto = {dc \<in> dc_crypto. inversion dc = x}" by auto
+    ultimately show "2 / 2^n = prob (inversion -` {x} \<inter> space (uniform_count_measure dc_crypto))"
+      using `0 < n`
+      by (simp add: card_inversion card_dc_crypto finite_dc_crypto
+                    subset_eq space_uniform_count_measure measure_uniform_count_measure)
+  qed
 
-  { have "\<H>(inversion|payer) =
-        - (\<Sum>x\<in>inversion`dc_crypto. (\<Sum>z\<in>Some ` {0..<n}. ?dIP {(x, z)} * log 2 (?dIP {(x, z)} / ?dP {z})))"
-      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)))"
-      unfolding neg_equal_iff_equal
-    proof (rule setsum_cong[OF refl], rule setsum_cong[OF refl])
-      fix x z assume x: "x \<in> inversion`dc_crypto" and z: "z \<in> Some ` {0..<n}"
-      hence "(\<lambda>x. (inversion x, payer x)) -` {(x, z)} \<inter> dc_crypto =
-          {dc \<in> dc_crypto. payer dc = Some (the z) \<and> inversion dc = x}"
-        by (auto simp add: payer_def)
-      moreover from x z obtain i where "z = Some i" and "i < n" by auto
-      moreover from x have "length x = n" by (auto simp: inversion_def [abs_def] dc_crypto)
-      ultimately
-      have "?dIP {(x, z)} = 2 / (real n * 2^n)" using x
-        by (auto simp add: card_dc_crypto card_payer_and_inversion distribution_def)
-      moreover
-      from z have "payer -` {z} \<inter> dc_crypto = {z} \<times> {xs. length xs = n}"
-        by (auto simp: dc_crypto payer_def)
-      hence "card (payer -` {z} \<inter> dc_crypto) = 2^n"
-        using card_lists_length_eq[where A="UNIV::bool set"]
-        by (simp add: card_cartesian_product_singleton)
-      hence "?dP {z} = 1 / real n"
-        by (simp add: distribution_def card_dc_crypto)
-      ultimately
-      show "?dIP {(x,z)} * log 2 (?dIP {(x,z)} / ?dP {z}) =
-       2 / (real n * 2^n) * (1 - real n)"
-        by (simp add: log_divide log_nat_power[of 2])
-    qed
-    also have "... = real n - 1"
-      using n finite_space
-      by (simp add: card_image_inversion card_image[OF inj_Some] field_simps real_eq_of_nat[symmetric])
-    finally have "\<H>(inversion|payer) = real n - 1" . }
-  moreover
-  { have "\<H>(inversion) = - (\<Sum>x \<in> inversion`dc_crypto. ?dI {x} * log 2 (?dI {x}))"
-      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])
-      fix x assume x_inv: "x \<in> inversion ` dc_crypto"
-      hence "length x = n" by (auto simp: inversion_def [abs_def] dc_crypto)
-      moreover have "inversion -` {x} \<inter> dc_crypto = {dc \<in> dc_crypto. inversion dc = x}" by auto
-      ultimately have "?dI {x} = 2 / 2^n" using `0 < n`
-        by (auto simp: card_inversion[OF x_inv] card_dc_crypto distribution_def)
-      thus "?dI {x} * log 2 (?dI {x}) = 2 * (1 - real n) / 2^n"
-        by (simp add: log_divide log_nat_power power_le_zero_eq inverse_eq_divide)
-    qed
-    also have "... = real n - 1"
-      by (simp add: card_image_inversion real_of_nat_def[symmetric] field_simps)
-    finally have "\<H>(inversion) = real n - 1" .
-  }
-  ultimately show ?thesis
-    unfolding mutual_information_eq_entropy_conditional_entropy[OF simple_function_finite simple_function_finite]
+  show "simple_distributed (uniform_count_measure dc_crypto) payer (\<lambda>x. 1 / real n)"
+  proof (rule simple_distributedI)
+    show "simple_function (uniform_count_measure dc_crypto) payer"
+      using finite_dc_crypto
+      by (auto simp: simple_function_def space_uniform_count_measure sets_uniform_count_measure)
+    fix z assume "z \<in> payer ` space (uniform_count_measure dc_crypto)"
+    then have "payer -` {z} \<inter> dc_crypto = {z} \<times> {xs. length xs = n}"
+      by (auto simp: dc_crypto payer_def space_uniform_count_measure)
+    hence "card (payer -` {z} \<inter> dc_crypto) = 2^n"
+      using card_lists_length_eq[where A="UNIV::bool set"]
+      by (simp add: card_cartesian_product_singleton)
+    then show "1 / real n = prob (payer -` {z} \<inter> space (uniform_count_measure dc_crypto))"
+      using finite_dc_crypto
+      by (subst measure_uniform_count_measure) (auto simp add: card_dc_crypto space_uniform_count_measure)
+  qed
+
+  show "simple_distributed (uniform_count_measure dc_crypto) (\<lambda>x. (inversion x, payer x)) (\<lambda>x. 2 / (real n *2^n))"
+  proof (rule simple_distributedI)
+    show "simple_function (uniform_count_measure dc_crypto) (\<lambda>x. (inversion x, payer x))"
+      using finite_dc_crypto
+      by (auto simp: simple_function_def space_uniform_count_measure sets_uniform_count_measure)
+    fix x assume "x \<in> (\<lambda>x. (inversion x, payer x)) ` space (uniform_count_measure dc_crypto)"
+    then obtain i xs where x: "x = (inversion (Some i, xs), payer (Some i, xs))"
+      and "i < n" "length xs = n"
+      by (simp add: image_iff space_uniform_count_measure dc_crypto Bex_def) blast
+    then have xs: "inversion (Some i, xs) \<in> inversion`dc_crypto" and i: "Some i \<in> Some ` {0..<n}"
+      and x: "x = (inversion (Some i, xs), Some i)" by (simp_all add: payer_def dc_crypto)
+    moreover def ys \<equiv> "inversion (Some i, xs)"
+    ultimately have ys: "ys \<in> inversion`dc_crypto"
+      and "Some i \<in> Some ` {0..<n}" "x = (ys, Some i)" by simp_all
+    then have "(\<lambda>x. (inversion x, payer x)) -` {x} \<inter> space (uniform_count_measure dc_crypto) =
+      {dc \<in> dc_crypto. payer dc = Some (the (Some i)) \<and> inversion dc = ys}"
+      by (auto simp add: payer_def space_uniform_count_measure)
+    then show "2 / (real n * 2 ^ n) = prob ((\<lambda>x. (inversion x, payer x)) -` {x} \<inter> space (uniform_count_measure dc_crypto))"
+      using `i < n` ys
+      by (simp add: measure_uniform_count_measure card_payer_and_inversion finite_dc_crypto subset_eq card_dc_crypto)
+  qed
+
+  show "\<forall>x\<in>space (uniform_count_measure dc_crypto). 2 / (real n * 2 ^ n) = 2 / 2 ^ n * (1 / real n) "
     by simp
 qed
 
--- a/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Mon Apr 23 12:23:23 2012 +0100
+++ b/src/HOL/Probability/ex/Koepf_Duermuth_Countermeasure.thy	Mon Apr 23 12:14:35 2012 +0200
@@ -103,49 +103,14 @@
       lessThan_Suc_eq_insert_0 setprod_reindex setsum_left_distrib[symmetric] setsum_right_distrib[symmetric])
 qed
 
-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])"
+declare space_point_measure[simp]
 
-  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]
+declare sets_point_measure[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
-
-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}))"
-
-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 information_space)
-  finite_guessing_entropy ("\<G>'(_')") where
-  "\<G>(X) \<equiv> guessing_entropy b X"
+lemma measure_point_measure:
+  "finite \<Omega> \<Longrightarrow> A \<subseteq> \<Omega> \<Longrightarrow> (\<And>x. x \<in> \<Omega> \<Longrightarrow> 0 \<le> p x) \<Longrightarrow>
+    measure (point_measure \<Omega> (\<lambda>x. ereal (p x))) A = (\<Sum>i\<in>A. p i)"
+  unfolding measure_def by (subst emeasure_point_measure_finite) auto
 
 locale finite_information =
   fixes \<Omega> :: "'a set"
@@ -159,17 +124,14 @@
 lemma (in finite_information) positive_p_sum[simp]: "0 \<le> setsum p X"
    by (auto intro!: setsum_nonneg)
 
-sublocale finite_information \<subseteq> finite_measure_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. ereal (setsum p S)\<rparr>"
-  by (rule finite_measure_spaceI) (simp_all add: setsum_Un_disjoint finite_subset)
+sublocale finite_information \<subseteq> prob_space "point_measure \<Omega> p"
+  by default (simp add: one_ereal_def emeasure_point_measure_finite)
 
-sublocale finite_information \<subseteq> finite_prob_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. ereal (setsum p S)\<rparr>"
-  by default (simp add: one_ereal_def)
-
-sublocale finite_information \<subseteq> information_space "\<lparr> space = \<Omega>, sets = Pow \<Omega>, measure = \<lambda>S. ereal (setsum p S)\<rparr>" b
+sublocale finite_information \<subseteq> information_space "point_measure \<Omega> p" b
   by default simp
 
-lemma (in finite_information) \<mu>'_eq: "A \<subseteq> \<Omega> \<Longrightarrow> \<mu>' A = setsum p A"
-  unfolding \<mu>'_def by auto
+lemma (in finite_information) \<mu>'_eq: "A \<subseteq> \<Omega> \<Longrightarrow> prob A = setsum p A"
+  by (auto simp: measure_point_measure)
 
 locale koepf_duermuth = K: finite_information keys K b + M: finite_information messages M b
     for b :: real
@@ -247,14 +209,19 @@
 abbreviation
  "p A \<equiv> setsum P A"
 
+abbreviation
+  "\<mu> \<equiv> point_measure msgs P"
+
 abbreviation probability ("\<P>'(_') _") where
- "\<P>(X) x \<equiv> distribution X x"
+  "\<P>(X) x \<equiv> measure \<mu> (X -` x \<inter> msgs)"
 
-abbreviation joint_probability ("\<P>'(_, _') _") where
- "\<P>(X, Y) x \<equiv> joint_distribution X Y x"
+abbreviation joint_probability ("\<P>'(_ ; _') _") where
+  "\<P>(X ; Y) x \<equiv> \<P>(\<lambda>x. (X x, Y x)) x"
 
-abbreviation conditional_probability ("\<P>'(_|_') _") where
- "\<P>(X|Y) x \<equiv> \<P>(X, Y) x / \<P>(Y) (snd`x)"
+no_notation disj (infixr "|" 30)
+
+abbreviation conditional_probability ("\<P>'(_ | _') _") where
+  "\<P>(X | Y) x \<equiv> (\<P>(X ; Y) x) / \<P>(Y) (snd`x)"
 
 notation
   entropy_Pow ("\<H>'( _ ')")
@@ -280,8 +247,8 @@
   from assms have *:
       "fst -` {k} \<inter> msgs = {k}\<times>{ms. set ms \<subseteq> messages \<and> length ms = n}"
     unfolding msgs_def by auto
-  show "\<P>(fst) {k} = K k"
-    apply (simp add: \<mu>'_eq distribution_def)
+  show "(\<P>(fst) {k}) = K k"
+    apply (simp add: \<mu>'_eq)
     apply (simp add: * P_def)
     apply (simp add: setsum_cartesian_product')
     using setprod_setsum_distrib_lists[OF M.finite_space, of M n "\<lambda>x x. True"] `k \<in> keys`
@@ -297,6 +264,67 @@
     unfolding msgs_def fst_image_times if_not_P[OF *] by simp
 qed
 
+lemma setsum_distribution_cut:
+  "\<P>(X) {x} = (\<Sum>y \<in> Y`space \<mu>. \<P>(X ; Y) {(x, y)})"
+  by (subst finite_measure_finite_Union[symmetric])
+     (auto simp add: disjoint_family_on_def inj_on_def
+           intro!: arg_cong[where f=prob])
+
+lemma prob_conj_imp1:
+  "prob ({x. Q x} \<inter> msgs) = 0 \<Longrightarrow> prob ({x. Pr x \<and> Q x} \<inter> msgs) = 0"
+  using finite_measure_mono[of "{x. Pr x \<and> Q x} \<inter> msgs" "{x. Q x} \<inter> msgs"]
+  using measure_nonneg[of \<mu> "{x. Pr x \<and> Q x} \<inter> msgs"]
+  by (simp add: subset_eq)
+
+lemma prob_conj_imp2:
+  "prob ({x. Pr x} \<inter> msgs) = 0 \<Longrightarrow> prob ({x. Pr x \<and> Q x} \<inter> msgs) = 0"
+  using finite_measure_mono[of "{x. Pr x \<and> Q x} \<inter> msgs" "{x. Pr x} \<inter> msgs"]
+  using measure_nonneg[of \<mu> "{x. Pr x \<and> Q x} \<inter> msgs"]
+  by (simp add: subset_eq)
+
+lemma simple_function_finite: "simple_function \<mu> f"
+  by (simp add: simple_function_def)
+
+lemma entropy_commute: "\<H>(\<lambda>x. (X x, Y x)) = \<H>(\<lambda>x. (Y x, X x))"
+  apply (subst (1 2) entropy_simple_distributed[OF simple_distributedI[OF simple_function_finite refl]])
+  unfolding space_point_measure
+proof -
+  have eq: "(\<lambda>x. (X x, Y x)) ` msgs = (\<lambda>(x, y). (y, x)) ` (\<lambda>x. (Y x, X x)) ` msgs"
+    by auto
+  show "- (\<Sum>x\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {x}) * log b (\<P>(X ; Y) {x})) =
+    - (\<Sum>x\<in>(\<lambda>x. (Y x, X x)) ` msgs. (\<P>(Y ; X) {x}) * log b (\<P>(Y ; X) {x}))"
+    unfolding eq
+    apply (subst setsum_reindex)
+    apply (auto simp: vimage_def inj_on_def intro!: setsum_cong arg_cong[where f="\<lambda>x. prob x * log b (prob x)"])
+    done
+qed
+
+lemma (in -) measure_eq_0I: "A = {} \<Longrightarrow> measure M A = 0" by simp
+
+lemma conditional_entropy_eq_ce_with_hypothesis:
+  "\<H>(X | Y) = -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) *
+     log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))"
+  apply (subst conditional_entropy_eq[OF
+    simple_distributedI[OF simple_function_finite refl]
+    simple_function_finite
+    simple_distributedI[OF simple_function_finite refl]])
+  unfolding space_point_measure
+proof -
+  have "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
+    - (\<Sum>x\<in>X`msgs. (\<Sum>y\<in>Y`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
+    unfolding setsum_cartesian_product
+    apply (intro arg_cong[where f=uminus] setsum_mono_zero_left)
+    apply (auto simp: vimage_def image_iff intro!: measure_eq_0I)
+    apply metis
+    done
+  also have "\<dots> = - (\<Sum>y\<in>Y`msgs. (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))))"
+    by (subst setsum_commute) rule
+  also have "\<dots> = -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))"
+    by (auto simp add: setsum_right_distrib vimage_def intro!: setsum_cong prob_conj_imp1)
+  finally show "- (\<Sum>(x, y)\<in>(\<lambda>x. (X x, Y x)) ` msgs. (\<P>(X ; Y) {(x, y)}) * log b ((\<P>(X ; Y) {(x, y)}) / (\<P>(Y) {y}))) =
+    -(\<Sum>y\<in>Y`msgs. (\<P>(Y) {y}) * (\<Sum>x\<in>X`msgs. (\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}) * log b ((\<P>(X ; Y) {(x,y)}) / (\<P>(Y) {y}))))" .
+qed
+
 lemma ce_OB_eq_ce_t: "\<H>(fst | OB) = \<H>(fst | t\<circ>OB)"
 proof -
   txt {* Lemma 2 *}
@@ -314,22 +342,22 @@
       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 =
+      have "(\<P>(OB ; fst) {(obs, k)}) / K k =
           p ({k}\<times>{ms. (k,ms) \<in> msgs \<and> OB (k,ms) = obs}) / K k"
-        apply (simp add: distribution_def \<mu>'_eq) by (auto intro!: arg_cong[where f=p])
+        apply (simp add: \<mu>'_eq) 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]) ..
-      finally have "\<P>(OB, fst) {(obs, k)} / K k =
+      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"
+    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)}"
+    then have "(\<P>(OB ; fst) {(obs, k)}) = (\<P>(OB ; fst) {(obs', k)})"
       using `K k \<noteq> 0` by auto }
   note t_eq_imp = this
 
@@ -339,13 +367,13 @@
       (\<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 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)}"
+    have "\<P>(t\<circ>OB ; fst) {(t obs, k)} = (\<Sum>obs'\<in>?S obs. \<P>(OB ; fst) {(obs', k)})"
+      unfolding comp_def
+      using finite_measure_finite_Union[OF _ df]
+      by (auto 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)}" .}
+    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"
@@ -359,11 +387,15 @@
     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 distribution_order(7,8)[where X=fst and x=k and Y="t\<circ>OB" and y="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_distribution_cut(2)[of "t\<circ>OB" fst "t obs", symmetric]
-      by (auto intro!: setsum_cong distribution_order(8))
+      using finite_measure_mono[of "{x. fst x = k \<and> t (OB x) = t obs} \<inter> msgs" "{x. fst x = k} \<inter> msgs"]
+      using measure_nonneg[of \<mu> "{x. fst x = k \<and> t (OB x) = t obs} \<inter> msgs"]
+      by (auto simp add: vimage_def conj_commute subset_eq)
+    also have "(\<P>(t\<circ>OB) {t obs}) = (\<Sum>k'\<in>keys. (\<P>(t\<circ>OB|fst) {(t obs, k')}) * (\<P>(fst) {k'}))"
+      using finite_measure_mono[of "{x. t (OB x) = t obs \<and> fst x = k} \<inter> msgs" "{x. fst x = k} \<inter> msgs"]
+      using measure_nonneg[of \<mu> "{x. fst x = k \<and> t (OB x) = t obs} \<inter> msgs"]
+      apply (simp add: setsum_distribution_cut[of "t\<circ>OB" "t obs" fst])
+      apply (auto intro!: setsum_cong simp: subset_eq vimage_def prob_conj_imp1)
+      done
     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`
@@ -371,10 +403,10 @@
           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_distribution_cut(2)[of OB fst obs, symmetric]
-      by (auto intro!: setsum_cong distribution_order(8))
+      using setsum_distribution_cut[of OB obs fst]
+      by (auto intro!: setsum_cong simp: prob_conj_imp1 vimage_def)
     also have "\<P>(OB | fst) {(obs, k)} * \<P>(fst) {k} / \<P>(OB) {obs} = \<P>(fst | OB) {(k, obs)}"
-      by (subst joint_distribution_commute) (auto intro!: distribution_order(8))
+      by (auto simp: vimage_def conj_commute prob_conj_imp2)
     finally have "\<P>(fst | t\<circ>OB) {(k, t obs)} = \<P>(fst | OB) {(k, obs)}" . }
   note CP_T_eq_CP_O = this
 
@@ -396,15 +428,15 @@
     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 finite_measure_finite_Union[OF _ _ df]
+      unfolding comp_def
+      using finite_measure_finite_Union[OF _ df]
       by (force simp add: * intro!: setsum_nonneg) }
   note P_t_sum_P_O = this
 
   txt {* Lemma 3 *}
+  txt {* Lemma 3 *}
   have "\<H>(fst | OB) = -(\<Sum>obs\<in>OB`msgs. \<P>(OB) {obs} * ?Ht (t obs))"
-    unfolding conditional_entropy_eq_ce_with_hypothesis[OF
-      simple_function_finite simple_function_finite] using * by simp
+    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)
@@ -418,8 +450,7 @@
     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[OF
-      simple_function_finite simple_function_finite]
+    unfolding conditional_entropy_eq_ce_with_hypothesis
     by (simp add: comp_def image_image[symmetric])
   finally show ?thesis .
 qed
@@ -433,11 +464,11 @@
     unfolding ce_OB_eq_ce_t ..
   also have "\<dots> = \<H>(t\<circ>OB) - \<H>(t\<circ>OB | fst)"
     unfolding entropy_chain_rule[symmetric, OF simple_function_finite simple_function_finite] sign_simps
-    by (subst entropy_commute[OF simple_function_finite simple_function_finite]) simp
+    by (subst entropy_commute) simp
   also have "\<dots> \<le> \<H>(t\<circ>OB)"
-    using conditional_entropy_positive[of "t\<circ>OB" fst] by simp
+    using conditional_entropy_nonneg[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
+    using entropy_le_card[of "t\<circ>OB", OF simple_distributedI[OF simple_function_finite refl]] 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)