fixed document generation for HOL-Probability
authorhoelzl
Mon, 19 May 2014 13:44:13 +0200
changeset 56994 8d5e5ec1cac3
parent 56993 e5366291d6aa
child 56995 61855ade6c7e
fixed document generation for HOL-Probability
src/HOL/Probability/Binary_Product_Measure.thy
src/HOL/Probability/Bochner_Integration.thy
src/HOL/Probability/Borel_Space.thy
src/HOL/Probability/Caratheodory.thy
src/HOL/Probability/Finite_Product_Measure.thy
src/HOL/Probability/Information.thy
src/HOL/Probability/Measure_Space.thy
src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy
src/HOL/Probability/Probability_Measure.thy
src/HOL/Probability/Radon_Nikodym.thy
src/HOL/Probability/Sigma_Algebra.thy
src/HOL/Probability/document/root.tex
src/HOL/Probability/ex/Dining_Cryptographers.thy
src/HOL/ROOT
--- a/src/HOL/Probability/Binary_Product_Measure.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Binary_Product_Measure.thy	Mon May 19 13:44:13 2014 +0200
@@ -14,7 +14,7 @@
 lemma rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
   by auto
 
-section "Binary products"
+subsection "Binary products"
 
 definition pair_measure (infixr "\<Otimes>\<^sub>M" 80) where
   "A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B)
@@ -480,7 +480,7 @@
     done
 qed
 
-section "Fubinis theorem"
+subsection "Fubinis theorem"
 
 lemma measurable_compose_Pair1:
   "x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^sub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L"
@@ -557,7 +557,7 @@
   shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M2) \<partial>M1)"
   unfolding positive_integral_snd[OF assms] M2.positive_integral_fst[OF assms] ..
 
-section {* Products on counting spaces, densities and distributions *}
+subsection {* Products on counting spaces, densities and distributions *}
 
 lemma sigma_sets_pair_measure_generator_finite:
   assumes "finite A" and "finite B"
--- a/src/HOL/Probability/Bochner_Integration.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Bochner_Integration.thy	Mon May 19 13:44:13 2014 +0200
@@ -1519,7 +1519,7 @@
     integral\<^sup>L M f \<le> integral\<^sup>L M g"
   by (intro integral_mono_AE) auto
 
-section {* Measure spaces with an associated density *}
+subsection {* Measure spaces with an associated density *}
 
 lemma integrable_density:
   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}" and g :: "'a \<Rightarrow> real"
@@ -1673,7 +1673,7 @@
     has_bochner_integral M (\<lambda>x. f (g x)) x \<Longrightarrow> has_bochner_integral (distr M N g) f x"
   by (simp add: has_bochner_integral_iff integrable_distr_eq integral_distr)
 
-section {* Lebesgue integration on @{const count_space} *}
+subsection {* Lebesgue integration on @{const count_space} *}
 
 lemma integrable_count_space:
   fixes f :: "'a \<Rightarrow> 'b::{banach,second_countable_topology}"
@@ -1703,7 +1703,7 @@
   by (subst lebesgue_integral_count_space_finite_support)
      (auto intro!: setsum_mono_zero_cong_left)
 
-section {* Point measure *}
+subsection {* Point measure *}
 
 lemma lebesgue_integral_point_measure_finite:
   fixes g :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
@@ -1720,7 +1720,7 @@
   apply (auto simp: AE_count_space integrable_count_space)
   done
 
-subsection {* Legacy lemmas for the real-valued Lebesgue integral\<^sup>L *}
+subsection {* Legacy lemmas for the real-valued Lebesgue integral *}
 
 lemma real_lebesgue_integral_def:
   assumes f: "integrable M f"
@@ -2080,7 +2080,7 @@
 qed
 
 
-section "Lebesgue integration on countable spaces"
+subsection {* Lebesgue integration on countable spaces *}
 
 lemma integral_on_countable:
   fixes f :: "real \<Rightarrow> real"
--- a/src/HOL/Probability/Borel_Space.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Borel_Space.thy	Mon May 19 13:44:13 2014 +0200
@@ -11,7 +11,7 @@
   "~~/src/HOL/Multivariate_Analysis/Multivariate_Analysis"
 begin
 
-section "Generic Borel spaces"
+subsection {* Generic Borel spaces *}
 
 definition borel :: "'a::topological_space measure" where
   "borel = sigma UNIV {S. open S}"
@@ -213,7 +213,7 @@
     unfolding eq by (rule borel_measurable_continuous_on[OF H]) auto
 qed
 
-section "Borel spaces on euclidean spaces"
+subsection {* Borel spaces on euclidean spaces *}
 
 lemma borel_measurable_inner[measurable (raw)]:
   fixes f g :: "'a \<Rightarrow> 'b::{second_countable_topology, real_inner}"
@@ -1140,7 +1140,7 @@
   shows "(\<lambda>x. (\<Sum>i. f i x)) \<in> borel_measurable M"
   unfolding suminf_def sums_def[abs_def] lim_def[symmetric] by simp
 
-section "LIMSEQ is borel measurable"
+subsection {* LIMSEQ is borel measurable *}
 
 lemma borel_measurable_LIMSEQ:
   fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
--- a/src/HOL/Probability/Caratheodory.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Caratheodory.thy	Mon May 19 13:44:13 2014 +0200
@@ -45,7 +45,7 @@
                      SUP_ereal_setsum[symmetric] incseq_setsumI setsum_nonneg)
 qed
 
-subsection {* Measure Spaces *}
+subsection {* Characterizations of Measures *}
 
 definition subadditive where "subadditive M f \<longleftrightarrow>
   (\<forall>x\<in>M. \<forall>y\<in>M. x \<inter> y = {} \<longrightarrow> f (x \<union> y) \<le> f x + f y)"
@@ -54,9 +54,6 @@
   (\<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 \<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"
 
@@ -67,7 +64,10 @@
   "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)
 
-subsection {* Lambda Systems *}
+subsubsection {* Lambda Systems *}
+
+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}"
 
 lemma (in algebra) lambda_system_eq:
   shows "lambda_system \<Omega> M f = {l \<in> M. \<forall>x \<in> M. f (x \<inter> l) + f (x - l) = f x}"
@@ -643,6 +643,8 @@
   by (simp add: measure_space_def positive_def countably_additive_def)
      blast
 
+subsection {* Caratheodory's theorem *}
+
 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> M. \<mu> s = f s) \<and> measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
@@ -670,8 +672,6 @@
     by (intro exI[of _ ?infm]) auto
 qed
 
-subsubsection {*Alternative instances of caratheodory*}
-
 lemma (in ring_of_sets) caratheodory_empty_continuous:
   assumes f: "positive M f" "additive M f" and fin: "\<And>A. A \<in> M \<Longrightarrow> f A \<noteq> \<infinity>"
   assumes cont: "\<And>A. range A \<subseteq> M \<Longrightarrow> decseq A \<Longrightarrow> (\<Inter>i. A i) = {} \<Longrightarrow> (\<lambda>i. f (A i)) ----> 0"
@@ -680,7 +680,7 @@
   show "\<forall>A\<in>M. f A \<noteq> \<infinity>" using fin by auto
 qed (rule cont)
 
-section {* Volumes *}
+subsection {* Volumes *}
 
 definition volume :: "'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> bool" where
   "volume M f \<longleftrightarrow>
@@ -818,7 +818,7 @@
   qed
 qed
 
-section {* Caratheodory on semirings *}
+subsubsection {* Caratheodory on semirings *}
 
 theorem (in semiring_of_sets) caratheodory:
   assumes pos: "positive M \<mu>" and ca: "countably_additive M \<mu>"
--- a/src/HOL/Probability/Finite_Product_Measure.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Finite_Product_Measure.thy	Mon May 19 13:44:13 2014 +0200
@@ -11,7 +11,7 @@
 lemma split_const: "(\<lambda>(i, j). c) = (\<lambda>_. c)"
   by auto
 
-subsubsection {* Merge two extensional functions *}
+subsubsection {* More about Function restricted by @{const extensional}  *}
 
 definition
   "merge I J = (\<lambda>(x, y) i. if i \<in> I then x i else if i \<in> J then y i else undefined)"
@@ -105,9 +105,9 @@
   "I \<inter> J = {} \<Longrightarrow> merge I J -` Pi\<^sub>E (I \<union> J) E = Pi I E \<times> Pi J E"
   by (auto simp: restrict_Pi_cancel PiE_def)
 
-section "Finite product spaces"
+subsection {* Finite product spaces *}
 
-section "Products"
+subsubsection {* Products *}
 
 definition prod_emb where
   "prod_emb I M K X = (\<lambda>x. restrict x K) -` X \<inter> (PIE i:I. space (M i))"
--- a/src/HOL/Probability/Information.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Information.thy	Mon May 19 13:44:13 2014 +0200
@@ -26,7 +26,7 @@
   "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and
   "(X = (A, B)) \<longleftrightarrow> (fst X = A \<and> snd X = B)" by auto
 
-section "Information theory"
+subsection "Information theory"
 
 locale information_space = prob_space +
   fixes b :: real assumes b_gt_1: "1 < b"
--- a/src/HOL/Probability/Measure_Space.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Measure_Space.thy	Mon May 19 13:44:13 2014 +0200
@@ -73,7 +73,7 @@
   represent sigma algebras (with an arbitrary emeasure).
 *}
 
-section "Extend binary sets"
+subsection "Extend binary sets"
 
 lemma LIMSEQ_binaryset:
   assumes f: "f {} = 0"
@@ -105,7 +105,7 @@
   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>} *}
+subsection {* Properties of a premeasure @{term \<mu>} *}
 
 text {*
   The definitions for @{const positive} and @{const countably_additive} should be here, by they are
@@ -429,7 +429,7 @@
   using empty_continuous_imp_continuous_from_below[OF f fin] cont
   by blast
 
-section {* Properties of @{const emeasure} *}
+subsection {* 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)
@@ -761,7 +761,7 @@
     by (simp add: emeasure_countably_additive)
 qed simp_all
 
-section "@{text \<mu>}-null sets"
+subsection {* @{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}"
@@ -853,7 +853,7 @@
     by (subst plus_emeasure[symmetric]) auto
 qed
 
-section "Formalize almost everywhere"
+subsection {* The almost everywhere filter (i.e.\ quantifier) *}
 
 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)"
@@ -1053,7 +1053,7 @@
   shows "emeasure M A = emeasure M B"
   using assms by (safe intro!: antisym emeasure_mono_AE) auto
 
-section {* @{text \<sigma>}-finite Measures *}
+subsection {* @{text \<sigma>}-finite Measures *}
 
 locale sigma_finite_measure =
   fixes M :: "'a measure"
@@ -1103,7 +1103,7 @@
   qed (force simp: incseq_def)+
 qed
 
-section {* Measure space induced by distribution of @{const measurable}-functions *}
+subsection {* 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))"
@@ -1187,7 +1187,7 @@
   by (auto simp add: emeasure_distr measurable_space
            intro!: arg_cong[where f="emeasure M"] measure_eqI)
 
-section {* Real measure values *}
+subsection {* 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)
@@ -1324,7 +1324,7 @@
     by (intro lim_real_of_ereal) simp
 qed
 
-section {* Measure spaces with @{term "emeasure M (space M) < \<infinity>"} *}
+subsection {* 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>"
@@ -1527,7 +1527,7 @@
   shows "measure M B = 0"
   using measure_space_inter[of B A] assms by (auto simp: ac_simps)
 
-section {* Counting space *}
+subsection {* Counting space *}
 
 lemma strict_monoI_Suc:
   assumes ord [simp]: "(\<And>n. f n < f (Suc n))" shows "strict_mono f"
@@ -1653,7 +1653,7 @@
   show "sigma_finite_measure (count_space A)" ..
 qed
 
-section {* Measure restricted to space *}
+subsection {* Measure restricted to space *}
 
 lemma emeasure_restrict_space:
   assumes "\<Omega> \<in> sets M" "A \<subseteq> \<Omega>"
--- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Mon May 19 13:44:13 2014 +0200
@@ -13,7 +13,7 @@
   "indicator A x \<le> (indicator B x::ereal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)"
   by (simp add: indicator_def not_le)
 
-section "Simple function"
+subsection "Simple function"
 
 text {*
 
@@ -501,7 +501,7 @@
   finally show "(\<lambda>x. f (T x)) -` {i} \<inter> space M \<in> sets M" .
 qed
 
-section "Simple integral"
+subsection "Simple integral"
 
 definition simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>S") where
   "integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
@@ -734,7 +734,7 @@
   then show ?thesis by simp
 qed
 
-section "Continuous positive integration"
+subsection {* Integral on nonnegative functions *}
 
 definition positive_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>P") where
   "integral\<^sup>P M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^sup>S M g)"
@@ -1556,7 +1556,9 @@
     by (simp add: F_def)
 qed
 
-section {* Distributions *}
+subsection {* Integral under concrete measures *}
+
+subsubsection {* Distributions *}
 
 lemma positive_integral_distr':
   assumes T: "T \<in> measurable M M'"
@@ -1587,7 +1589,7 @@
   by (subst (1 2) positive_integral_max_0[symmetric])
      (simp add: positive_integral_distr')
 
-section {* Lebesgue integration on @{const count_space} *}
+subsubsection {* Counting space *}
 
 lemma simple_function_count_space[simp]:
   "simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)"
@@ -1666,7 +1668,7 @@
   finally show ?thesis .
 qed
 
-section {* Measures with Restricted Space *}
+subsubsection {* Measures with Restricted Space *}
 
 lemma positive_integral_restrict_space:
   assumes \<Omega>: "\<Omega> \<in> sets M" and f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" "\<And>x. x \<in> space M - \<Omega> \<Longrightarrow> f x = 0"
@@ -1696,7 +1698,7 @@
     by (auto simp add: SUP_eq_iff measurable_restrict_space1 \<Omega> positive_integral_monotone_convergence_SUP)
 qed
 
-section {* Measure spaces with an associated density *}
+subsubsection {* 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>\<^sup>+ x. f x * indicator A x \<partial>M)"
@@ -1918,7 +1920,7 @@
     using f g by (subst density_density_eq) auto
 qed
 
-section {* Point measure *}
+subsubsection {* Point measure *}
 
 definition point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
   "point_measure A f = density (count_space A) f"
@@ -1991,7 +1993,7 @@
     integral\<^sup>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)
 
-section {* Uniform measure *}
+subsubsection {* Uniform measure *}
 
 definition "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)"
 
@@ -2019,7 +2021,7 @@
   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 *}
+subsubsection {* Uniform count measure *}
 
 definition "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)"
  
--- a/src/HOL/Probability/Probability_Measure.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy	Mon May 19 13:44:13 2014 +0200
@@ -349,7 +349,7 @@
   finally show ?thesis by simp
 qed
 
-section {* Distributions *}
+subsection {* Distributions *}
 
 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"
--- a/src/HOL/Probability/Radon_Nikodym.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Mon May 19 13:44:13 2014 +0200
@@ -769,7 +769,7 @@
     by (auto intro!: bexI[of _ "\<lambda>x. h x * f x"] simp: density_density_eq)
 qed
 
-section "Uniqueness of densities"
+subsection {* Uniqueness of densities *}
 
 lemma finite_density_unique:
   assumes borel: "f \<in> borel_measurable M" "g \<in> borel_measurable M"
@@ -1060,7 +1060,7 @@
   apply (auto simp: max_def intro!: measurable_If)
   done
 
-section "Radon-Nikodym derivative"
+subsection {* Radon-Nikodym derivative *}
 
 definition RN_deriv :: "'a measure \<Rightarrow> 'a measure \<Rightarrow> 'a \<Rightarrow> ereal" where
   "RN_deriv M N =
--- a/src/HOL/Probability/Sigma_Algebra.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/Sigma_Algebra.thy	Mon May 19 13:44:13 2014 +0200
@@ -5,7 +5,7 @@
     translated by Lawrence Paulson.
 *)
 
-header {* Sigma Algebras *}
+header {* Describing measurable sets *}
 
 theory Sigma_Algebra
 imports
@@ -33,9 +33,7 @@
 lemma (in subset_class) sets_into_space: "x \<in> M \<Longrightarrow> x \<subseteq> \<Omega>"
   by (metis PowD contra_subsetD space_closed)
 
-subsection {* Semiring of sets *}
-
-subsubsection {* Disjoint sets *}
+subsubsection {* Semiring of sets *}
 
 definition "disjoint A \<longleftrightarrow> (\<forall>a\<in>A. \<forall>b\<in>A. a \<noteq> b \<longrightarrow> a \<inter> b = {})"
 
@@ -255,7 +253,7 @@
   "X \<subseteq> S \<Longrightarrow> algebra S { {}, X, S - X, S }"
   by (auto simp: algebra_iff_Int)
 
-subsection {* Restricted algebras *}
+subsubsection {* Restricted algebras *}
 
 abbreviation (in algebra)
   "restricted_space A \<equiv> (op \<inter> A) ` M"
@@ -264,7 +262,7 @@
   assumes "A \<in> M" shows "algebra A (restricted_space A)"
   using assms by (auto simp: algebra_iff_Int)
 
-subsection {* Sigma Algebras *}
+subsubsection {* Sigma Algebras *}
 
 locale sigma_algebra = algebra +
   assumes countable_nat_UN [intro]: "\<And>A. range A \<subseteq> M \<Longrightarrow> (\<Union>i::nat. A i) \<in> M"
@@ -446,7 +444,7 @@
   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 *}
+subsubsection {* Binary Unions *}
 
 definition binary :: "'a \<Rightarrow> 'a \<Rightarrow> nat \<Rightarrow> 'a"
   where "binary a b =  (\<lambda>x. b)(0 := a)"
@@ -468,7 +466,7 @@
   by (auto simp add: range_binary_eq sigma_algebra_def sigma_algebra_axioms_def
          algebra_iff_Un Un_range_binary)
 
-subsection {* Initial Sigma Algebra *}
+subsubsection {* Initial Sigma Algebra *}
 
 text {*Sigma algebras can naturally be created as the closure of any set of
   M with regard to the properties just postulated.  *}
@@ -775,7 +773,7 @@
   qed
 qed
 
-subsection "Disjoint families"
+subsubsection "Disjoint families"
 
 definition
   disjoint_family_on  where
@@ -934,7 +932,7 @@
     by (intro disjointD[OF d]) auto
 qed
 
-subsection {* Ring generated by a semiring *}
+subsubsection {* Ring generated by a semiring *}
 
 definition (in semiring_of_sets)
   "generated_ring = { \<Union>C | C. C \<subseteq> M \<and> finite C \<and> disjoint C }"
@@ -1064,759 +1062,7 @@
     by (blast intro!: sigma_sets_mono elim: generated_ringE)
 qed (auto intro!: generated_ringI_Basic sigma_sets_mono)
 
-subsection {* 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 'a measure = "{(\<Omega>::'a set, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu> }"
-proof
-  have "sigma_algebra UNIV {{}, UNIV}"
-    by (auto simp: sigma_algebra_iff2)
-  then show "(UNIV, {{}, UNIV}, \<lambda>A. 0) \<in> {(\<Omega>, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu>} "
-    by (auto simp: measure_space_def positive_def countably_additive_def)
-qed
-
-definition space :: "'a measure \<Rightarrow> 'a set" where
-  "space M = fst (Rep_measure M)"
-
-definition sets :: "'a measure \<Rightarrow> 'a set set" where
-  "sets M = fst (snd (Rep_measure M))"
-
-definition emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> 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 sets!: sigma_algebra "space M" "sets M" for M :: "'a measure"
-  using measure_space[of M] by (auto simp: measure_space_def)
-
-definition measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
-  "measure_of \<Omega> A \<mu> = Abs_measure (\<Omega>, if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>},
-    \<lambda>a. if a \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> a else 0)"
-
-abbreviation "sigma \<Omega> A \<equiv> measure_of \<Omega> A (\<lambda>x. 0)"
-
-lemma measure_space_0: "A \<subseteq> Pow \<Omega> \<Longrightarrow> measure_space \<Omega> (sigma_sets \<Omega> A) (\<lambda>x. 0)"
-  unfolding measure_space_def
-  by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def)
-
-lemma sigma_algebra_trivial: "sigma_algebra \<Omega> {{}, \<Omega>}"
-by unfold_locales(fastforce intro: exI[where x="{{}}"] exI[where x="{\<Omega>}"])+
-
-lemma measure_space_0': "measure_space \<Omega> {{}, \<Omega>} (\<lambda>x. 0)"
-by(simp add: measure_space_def positive_def countably_additive_def sigma_algebra_trivial)
-
-lemma measure_space_closed:
-  assumes "measure_space \<Omega> M \<mu>"
-  shows "M \<subseteq> Pow \<Omega>"
-proof -
-  interpret sigma_algebra \<Omega> M using assms by(simp add: measure_space_def)
-  show ?thesis by(rule space_closed)
-qed
-
-lemma (in ring_of_sets) positive_cong_eq:
-  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> positive M \<mu>' = positive M \<mu>"
-  by (auto simp add: positive_def)
-
-lemma (in sigma_algebra) countably_additive_eq:
-  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> countably_additive M \<mu>' = countably_additive M \<mu>"
-  unfolding countably_additive_def
-  by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq)
-
-lemma measure_space_eq:
-  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a"
-  shows "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
-proof -
-  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" using closed by (rule sigma_algebra_sigma_sets)
-  from positive_cong_eq[OF eq, of "\<lambda>i. i"] countably_additive_eq[OF eq, of "\<lambda>i. i"] show ?thesis
-    by (auto simp: measure_space_def)
-qed
-
-lemma measure_of_eq:
-  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "(\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a)"
-  shows "measure_of \<Omega> A \<mu> = measure_of \<Omega> A \<mu>'"
-proof -
-  have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
-    using assms by (rule measure_space_eq)
-  with eq show ?thesis
-    by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure])
-qed
-
-lemma
-  shows space_measure_of_conv: "space (measure_of \<Omega> A \<mu>) = \<Omega>" (is ?space)
-  and sets_measure_of_conv:
-  "sets (measure_of \<Omega> A \<mu>) = (if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>})" (is ?sets)
-  and emeasure_measure_of_conv: 
-  "emeasure (measure_of \<Omega> A \<mu>) = 
-  (\<lambda>B. if B \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> B else 0)" (is ?emeasure)
-proof -
-  have "?space \<and> ?sets \<and> ?emeasure"
-  proof(cases "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>")
-    case True
-    from measure_space_closed[OF this] sigma_sets_superset_generator[of A \<Omega>]
-    have "A \<subseteq> Pow \<Omega>" by simp
-    hence "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A)
-      (\<lambda>a. if a \<in> sigma_sets \<Omega> A then \<mu> a else 0)"
-      by(rule measure_space_eq) auto
-    with True `A \<subseteq> Pow \<Omega>` show ?thesis
-      by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse)
-  next
-    case False thus ?thesis
-      by(cases "A \<subseteq> Pow \<Omega>")(simp_all add: Abs_measure_inverse measure_of_def sets_def space_def emeasure_def measure_space_0 measure_space_0')
-  qed
-  thus ?space ?sets ?emeasure by simp_all
-qed
-
-lemma [simp]:
-  assumes A: "A \<subseteq> Pow \<Omega>"
-  shows sets_measure_of: "sets (measure_of \<Omega> A \<mu>) = sigma_sets \<Omega> A"
-    and space_measure_of: "space (measure_of \<Omega> A \<mu>) = \<Omega>"
-using assms
-by(simp_all add: sets_measure_of_conv space_measure_of_conv)
-
-lemma (in sigma_algebra) sets_measure_of_eq[simp]: "sets (measure_of \<Omega> M \<mu>) = M"
-  using space_closed by (auto intro!: sigma_sets_eq)
-
-lemma (in sigma_algebra) space_measure_of_eq[simp]: "space (measure_of \<Omega> M \<mu>) = \<Omega>"
-  by (rule space_measure_of_conv)
-
-lemma measure_of_subset: "M \<subseteq> Pow \<Omega> \<Longrightarrow> M' \<subseteq> M \<Longrightarrow> sets (measure_of \<Omega> M' \<mu>) \<subseteq> sets (measure_of \<Omega> M \<mu>')"
-  by (auto intro!: sigma_sets_subseteq)
-
-lemma sigma_sets_mono'':
-  assumes "A \<in> sigma_sets C D"
-  assumes "B \<subseteq> D"
-  assumes "D \<subseteq> Pow C"
-  shows "sigma_sets A B \<subseteq> sigma_sets C D"
-proof
-  fix x assume "x \<in> sigma_sets A B"
-  thus "x \<in> sigma_sets C D"
-  proof induct
-    case (Basic a) with assms have "a \<in> D" by auto
-    thus ?case ..
-  next
-    case Empty show ?case by (rule sigma_sets.Empty)
-  next
-    from assms have "A \<in> sets (sigma C D)" by (subst sets_measure_of[OF `D \<subseteq> Pow C`])
-    moreover case (Compl a) hence "a \<in> sets (sigma C D)" by (subst sets_measure_of[OF `D \<subseteq> Pow C`])
-    ultimately have "A - a \<in> sets (sigma C D)" ..
-    thus ?case by (subst (asm) sets_measure_of[OF `D \<subseteq> Pow C`])
-  next
-    case (Union a)
-    thus ?case by (intro sigma_sets.Union)
-  qed
-qed
-
-lemma in_measure_of[intro, simp]: "M \<subseteq> Pow \<Omega> \<Longrightarrow> A \<in> M \<Longrightarrow> A \<in> sets (measure_of \<Omega> M \<mu>)"
-  by auto
-
-subsection {* 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)
-  thus ?thesis using X ms
-    by(simp add: M emeasure_measure_of_conv sets_measure_of_conv)
-qed
-
-lemma emeasure_measure_of_sigma:
-  assumes ms: "sigma_algebra \<Omega> M" "positive M \<mu>" "countably_additive M \<mu>"
-  assumes A: "A \<in> M"
-  shows "emeasure (measure_of \<Omega> M \<mu>) A = \<mu> A"
-proof -
-  interpret sigma_algebra \<Omega> M by fact
-  have "measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
-    using ms sigma_sets_eq by (simp add: measure_space_def)
-  thus ?thesis by(simp add: emeasure_measure_of_conv A)
-qed
-
-lemma measure_cases[cases type: measure]:
-  obtains (measure) \<Omega> A \<mu> where "x = Abs_measure (\<Omega>, A, \<mu>)" "\<forall>a\<in>-A. \<mu> a = 0" "measure_space \<Omega> A \<mu>"
-  by atomize_elim (cases x, auto)
-
-lemma sets_eq_imp_space_eq:
-  "sets M = sets M' \<Longrightarrow> space M = space M'"
-  using sets.top[of M] sets.top[of M'] sets.space_closed[of M] sets.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 emeasure_neq_0_sets: "emeasure M A \<noteq> 0 \<Longrightarrow> A \<in> sets M"
-  using emeasure_notin_sets[of A M] by blast
-
-lemma measure_notin_sets: "A \<notin> sets M \<Longrightarrow> measure M A = 0"
-  by (simp add: measure_def emeasure_notin_sets)
-
-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 AA': "A = A'" by simp
-  moreover from M.top N.top M.space_closed N.space_closed AA' have "\<Omega> = \<Omega>'" by auto
-  moreover { fix B have "\<mu> B = \<mu>' B"
-    proof cases
-      assume "B \<in> A"
-      with eq `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)
-
-subsection {* 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_sets_Collect:
-  assumes f: "f \<in> measurable M N" and P: "{x\<in>space N. P x} \<in> sets N" shows "{x\<in>space M. P (f x)} \<in> sets M"
-proof -
-  have "f -` {x \<in> space N. P x} \<inter> space M = {x\<in>space M. P (f x)}"
-    using measurable_space[OF f] by auto
-  with measurable_sets[OF f P] show ?thesis
-    by simp
-qed
-
-lemma measurable_sigma_sets:
-  assumes B: "sets N = sigma_sets \<Omega> A" "A \<subseteq> Pow \<Omega>"
-      and f: "f \<in> space M \<rightarrow> \<Omega>"
-      and ba: "\<And>y. y \<in> A \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
-  shows "f \<in> measurable M N"
-proof -
-  interpret A: sigma_algebra \<Omega> "sigma_sets \<Omega> A" using B(2) by (rule sigma_algebra_sigma_sets)
-  from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \<Omega>: "\<Omega> = space N" by force
-  
-  { fix X assume "X \<in> sigma_sets \<Omega> A"
-    then have "f -` X \<inter> space M \<in> sets M \<and> X \<subseteq> \<Omega>"
-      proof induct
-        case (Basic a) then show ?case
-          by (auto simp add: ba) (metis B(2) subsetD PowD)
-      next
-        case (Compl a)
-        have [simp]: "f -` \<Omega> \<inter> space M = space M"
-          by (auto simp add: funcset_mem [OF f])
-        then show ?case
-          by (auto simp add: vimage_Diff Diff_Int_distrib2 sets.compl_sets Compl)
-      next
-        case (Union a)
-        then show ?case
-          by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast
-      qed auto }
-  with f show ?thesis
-    by (auto simp add: measurable_def B \<Omega>)
-qed
-
-lemma measurable_measure_of:
-  assumes B: "N \<subseteq> Pow \<Omega>"
-      and f: "f \<in> space M \<rightarrow> \<Omega>"
-      and ba: "\<And>y. y \<in> N \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
-  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>)"
-proof -
-  have "sets (measure_of \<Omega> N \<mu>) = sigma_sets \<Omega> N"
-    using B by (rule sets_measure_of)
-  from this assms show ?thesis by (rule measurable_sigma_sets)
-qed
-
-lemma measurable_iff_measure_of:
-  assumes "N \<subseteq> Pow \<Omega>" "f \<in> space M \<rightarrow> \<Omega>"
-  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>) \<longleftrightarrow> (\<forall>A\<in>N. f -` A \<inter> space M \<in> sets M)"
-  by (metis assms in_measure_of measurable_measure_of assms measurable_sets)
-
-lemma measurable_cong_sets:
-  assumes sets: "sets M = sets M'" "sets N = sets N'"
-  shows "measurable M N = measurable M' N'"
-  using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def)
-
-lemma measurable_cong:
-  assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
-  shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
-  unfolding measurable_def using assms
-  by (simp cong: vimage_inter_cong Pi_cong)
-
-lemma measurable_cong_strong:
-  "M = N \<Longrightarrow> M' = N' \<Longrightarrow> (\<And>w. w \<in> space M \<Longrightarrow> f w = g w) \<Longrightarrow>
-    f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable N N'"
-  by (metis measurable_cong)
-
-lemma measurable_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_compose:
-  assumes f: "f \<in> measurable M N" and g: "g \<in> measurable N L"
-  shows "(\<lambda>x. g (f x)) \<in> measurable M L"
-proof -
-  have "\<And>A. (\<lambda>x. g (f x)) -` A \<inter> space M = f -` (g -` A \<inter> space N) \<inter> space M"
-    using measurable_space[OF f] by auto
-  with measurable_space[OF f] measurable_space[OF g] show ?thesis
-    by (auto intro: measurable_sets[OF f] measurable_sets[OF g]
-             simp del: vimage_Int simp add: measurable_def)
-qed
-
-lemma measurable_comp:
-  "f \<in> measurable M N \<Longrightarrow> g \<in> measurable N L \<Longrightarrow> g \<circ> f \<in> measurable M L"
-  using measurable_compose[of f M N g L] by (simp add: comp_def)
-
-lemma measurable_const:
-  "c \<in> space M' \<Longrightarrow> (\<lambda>x. c) \<in> measurable M M'"
-  by (auto simp add: measurable_def)
-
-lemma measurable_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!: sets.Un)
-qed
-
-lemma measurable_If_set:
-  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
-  assumes P: "A \<inter> space M \<in> sets M"
-  shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"
-proof (rule measurable_If[OF measure])
-  have "{x \<in> space M. x \<in> A} = A \<inter> space M" by auto
-  thus "{x \<in> space M. x \<in> A} \<in> sets M" using `A \<inter> space M \<in> sets M` by auto
-qed
-
-lemma measurable_ident: "id \<in> measurable M M"
-  by (auto simp add: measurable_def)
-
-lemma measurable_ident_sets:
-  assumes eq: "sets M = sets M'" shows "(\<lambda>x. x) \<in> measurable M M'"
-  using measurable_ident[of M]
-  unfolding id_def measurable_def eq sets_eq_imp_space_eq[OF eq] .
-
-lemma sets_Least:
-  assumes meas: "\<And>i::nat. {x\<in>space M. P i x} \<in> M"
-  shows "(\<lambda>x. LEAST j. P j x) -` A \<inter> space M \<in> sets M"
-proof -
-  { fix i have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M \<in> sets M"
-    proof cases
-      assume i: "(LEAST j. False) = i"
-      have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
-        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x})) \<union> (space M - (\<Union>i. {x\<in>space M. P i x}))"
-        by (simp add: set_eq_iff, safe)
-           (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality)
-      with meas show ?thesis
-        by (auto intro!: sets.Int)
-    next
-      assume i: "(LEAST j. False) \<noteq> i"
-      then have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
-        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x}))"
-      proof (simp add: set_eq_iff, safe)
-        fix x assume neq: "(LEAST j. False) \<noteq> (LEAST j. P j x)"
-        have "\<exists>j. P j x"
-          by (rule ccontr) (insert neq, auto)
-        then show "P (LEAST j. P j x) x" by (rule LeastI_ex)
-      qed (auto dest: Least_le intro!: Least_equality)
-      with meas show ?thesis
-        by auto
-    qed }
-  then have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) \<in> sets M"
-    by (intro sets.countable_UN) auto
-  moreover have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) =
-    (\<lambda>x. LEAST j. P j x) -` A \<inter> space M" by auto
-  ultimately show ?thesis by auto
-qed
-
-lemma measurable_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. (g \<circ> f) -` 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_comp)
-    apply (metis funcset_mem fab g)
-    apply (subst eq)
-    apply (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 {* 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:
-  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!: sets.finite_UN simp del: UN_simps) }
-  then show ?thesis
-    unfolding measurable_def by auto
-qed
-
-lemma measurable_count_space_eq2_countable:
-  fixes f :: "'a => 'c::countable"
-  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
-proof -
-  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
-    assume *: "\<And>a. a\<in>A \<Longrightarrow> f -` {a} \<inter> space M \<in> sets M"
-    have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)"
-      by auto
-    also have "\<dots> \<in> sets M"
-      using * `X \<subseteq> A` by (intro sets.countable_UN) auto
-    finally have "f -` X \<inter> space M \<in> sets M" . }
-  then show ?thesis
-    unfolding measurable_def by auto
-qed
-
-lemma measurable_compose_countable:
-  assumes f: "\<And>i::'i::countable. (\<lambda>x. f i x) \<in> measurable M N" and g: "g \<in> measurable M (count_space UNIV)"
-  shows "(\<lambda>x. f (g x) x) \<in> measurable M N"
-  unfolding measurable_def
-proof safe
-  fix x assume "x \<in> space M" then show "f (g x) x \<in> space N"
-    using f[THEN measurable_space] g[THEN measurable_space] by auto
-next
-  fix A assume A: "A \<in> sets N"
-  have "(\<lambda>x. f (g x) x) -` A \<inter> space M = (\<Union>i. (g -` {i} \<inter> space M) \<inter> (f i -` A \<inter> space M))"
-    by auto
-  also have "\<dots> \<in> sets M" using f[THEN measurable_sets, OF A] g[THEN measurable_sets]
-    by (auto intro!: sets.countable_UN measurable_sets)
-  finally show "(\<lambda>x. f (g x) x) -` A \<inter> space M \<in> sets M" .
-qed
-
-lemma measurable_count_space_const:
-  "(\<lambda>x. c) \<in> measurable M (count_space UNIV)"
-  by (simp add: measurable_const)
-
-lemma measurable_count_space:
-  "f \<in> measurable (count_space A) (count_space UNIV)"
-  by simp
-
-lemma measurable_compose_rev:
-  assumes f: "f \<in> measurable L N" and g: "g \<in> measurable M L"
-  shows "(\<lambda>x. f (g x)) \<in> measurable M N"
-  using measurable_compose[OF g f] .
-
-lemma measurable_count_space_eq_countable:
-  assumes "countable A"
-  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
-proof -
-  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
-    with `countable A` have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "countable X"
-      by (auto dest: countable_subset)
-    moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
-    ultimately have "f -` X \<inter> space M \<in> sets M"
-      using `X \<subseteq> A` by (auto intro!: sets.countable_UN' simp del: UN_simps) }
-  then show ?thesis
-    unfolding measurable_def by auto
-qed
-
-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
-  "vimage_algebra M S X = sigma S ((\<lambda>A. X -` A \<inter> S) ` sets M)"
-
-lemma sigma_algebra_preimages:
-  fixes f :: "'x \<Rightarrow> 'a"
-  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 A assume "A \<in> sets M"
-  moreover have "S - ?F A = ?F (space M - A)"
-    using assms by auto
-  ultimately show "S - ?F A \<in> ?F ` sets M"
-    by blast
-next
-  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 "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> M" "\<And>i. A i = ?F (b i)"
-    by auto
-  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 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 "f \<in> measurable (vimage_algebra M S f) M"
-  unfolding measurable_def using assms by force
-
-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 M S f) M2"
-proof -
-  note measurable_vimage_algebra[OF assms(2)]
-  from measurable_comp[OF this assms(1)]
-  show ?thesis by (simp add: comp_def)
-qed
-
-lemma sigma_sets_vimage:
-  assumes "f \<in> S' \<rightarrow> S" and "A \<subseteq> Pow S"
-  shows "sigma_sets S' ((\<lambda>X. f -` X \<inter> S') ` A) = (\<lambda>X. f -` X \<inter> S') ` sigma_sets S A"
-proof (intro set_eqI iffI)
-  let ?F = "\<lambda>X. f -` X \<inter> S'"
-  fix X assume "X \<in> sigma_sets S' (?F ` A)"
-  then show "X \<in> ?F ` sigma_sets S A"
-  proof induct
-    case (Basic X) then obtain X' where "X = ?F X'" "X' \<in> A"
-      by auto
-    then show ?case by auto
-  next
-    case Empty then show ?case
-      by (auto intro!: image_eqI[of _ _ "{}"] sigma_sets.Empty)
-  next
-    case (Compl X) then obtain X' where X: "X = ?F X'" and "X' \<in> sigma_sets S A"
-      by auto
-    then have "S - X' \<in> sigma_sets S A"
-      by (auto intro!: sigma_sets.Compl)
-    then show ?case
-      using X assms by (auto intro!: image_eqI[where x="S - X'"])
-  next
-    case (Union F)
-    then have "\<forall>i. \<exists>F'.  F' \<in> sigma_sets S A \<and> F i = f -` F' \<inter> S'"
-      by (auto simp: image_iff Bex_def)
-    from choice[OF this] obtain F' where
-      "\<And>i. F' i \<in> sigma_sets S A" and "\<And>i. F i = f -` F' i \<inter> S'"
-      by auto
-    then show ?case
-      by (auto intro!: sigma_sets.Union image_eqI[where x="\<Union>i. F' i"])
-  qed
-next
-  let ?F = "\<lambda>X. f -` X \<inter> S'"
-  fix X assume "X \<in> ?F ` sigma_sets S A"
-  then obtain X' where "X' \<in> sigma_sets S A" "X = ?F X'" by auto
-  then show "X \<in> sigma_sets S' (?F ` A)"
-  proof (induct arbitrary: X)
-    case Empty then show ?case by (auto intro: sigma_sets.Empty)
-  next
-    case (Compl X')
-    have "S' - (S' - X) \<in> sigma_sets S' (?F ` A)"
-      apply (rule sigma_sets.Compl)
-      using assms by (auto intro!: Compl.hyps simp: Compl.prems)
-    also have "S' - (S' - X) = X"
-      using assms Compl by auto
-    finally show ?case .
-  next
-    case (Union F)
-    have "(\<Union>i. f -` F i \<inter> S') \<in> sigma_sets S' (?F ` A)"
-      by (intro sigma_sets.Union Union.hyps) simp
-    also have "(\<Union>i. f -` F i \<inter> S') = X"
-      using assms Union by auto
-    finally show ?case .
-  qed auto
-qed
-
-subsection {* Restricted Space Sigma Algebra *}
-
-definition "restrict_space M \<Omega> = measure_of \<Omega> ((op \<inter> \<Omega>) ` sets M) (\<lambda>A. emeasure M (A \<inter> \<Omega>))"
-
-lemma space_restrict_space: "space (restrict_space M \<Omega>) = \<Omega>"
-  unfolding restrict_space_def by (subst space_measure_of) auto
-
-lemma sets_restrict_space: "\<Omega> \<subseteq> space M \<Longrightarrow> sets (restrict_space M \<Omega>) = (op \<inter> \<Omega>) ` sets M"
-  using sigma_sets_vimage[of "\<lambda>x. x" \<Omega> "space M" "sets M"]
-  unfolding restrict_space_def
-  by (subst sets_measure_of) (auto simp: sets.sigma_sets_eq Int_commute[of _ \<Omega>] sets.space_closed)
-
-lemma sets_restrict_space_iff:
-  "\<Omega> \<in> sets M \<Longrightarrow> A \<in> sets (restrict_space M \<Omega>) \<longleftrightarrow> (A \<subseteq> \<Omega> \<and> A \<in> sets M)"
-  by (subst sets_restrict_space) (auto dest: sets.sets_into_space)
-
-lemma measurable_restrict_space1:
-  assumes \<Omega>: "\<Omega> \<in> sets M" and f: "f \<in> measurable M N" shows "f \<in> measurable (restrict_space M \<Omega>) N"
-  unfolding measurable_def
-proof (intro CollectI conjI ballI)
-  show sp: "f \<in> space (restrict_space M \<Omega>) \<rightarrow> space N"
-    using measurable_space[OF f] sets.sets_into_space[OF \<Omega>] by (auto simp: space_restrict_space)
-
-  fix A assume "A \<in> sets N"
-  have "f -` A \<inter> space (restrict_space M \<Omega>) = (f -` A \<inter> space M) \<inter> \<Omega>"
-    using sets.sets_into_space[OF \<Omega>] by (auto simp: space_restrict_space)
-  also have "\<dots> \<in> sets (restrict_space M \<Omega>)"
-    unfolding sets_restrict_space_iff[OF \<Omega>]
-    using measurable_sets[OF f `A \<in> sets N`] \<Omega> by blast
-  finally show "f -` A \<inter> space (restrict_space M \<Omega>) \<in> sets (restrict_space M \<Omega>)" .
-qed
-
-lemma measurable_restrict_space2:
-  "\<Omega> \<in> sets N \<Longrightarrow> f \<in> space M \<rightarrow> \<Omega> \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> f \<in> measurable M (restrict_space N \<Omega>)"
-  by (simp add: measurable_def space_restrict_space sets_restrict_space_iff)
-
-subsection {* A Two-Element Series *}
+subsubsection {* A Two-Element Series *}
 
 definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "
   where "binaryset A B = (\<lambda>x. {})(0 := A, Suc 0 := B)"
@@ -1830,7 +1076,7 @@
 lemma UN_binaryset_eq: "(\<Union>i. binaryset A B i) = A \<union> B"
   by (simp add: SUP_def range_binaryset_eq)
 
-section {* Closed CDI *}
+subsubsection {* Closed CDI *}
 
 definition closed_cdi where
   "closed_cdi \<Omega> M \<longleftrightarrow>
@@ -2064,7 +1310,7 @@
     by blast
 qed
 
-subsection {* Dynkin systems *}
+subsubsection {* Dynkin systems *}
 
 locale dynkin_system = subset_class +
   assumes space: "\<Omega> \<in> M"
@@ -2126,7 +1372,7 @@
   show ?thesis using sets_into_space by (fastforce intro!: dynkin_systemI)
 qed
 
-subsection "Intersection stable algebras"
+subsubsection "Intersection sets systems"
 
 definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> M. \<forall> b \<in> M. a \<inter> b \<in> M)"
 
@@ -2162,7 +1408,7 @@
   qed auto
 qed
 
-subsection "Smallest Dynkin systems"
+subsubsection "Smallest Dynkin systems"
 
 definition dynkin where
   "dynkin \<Omega> M =  (\<Inter>{D. dynkin_system \<Omega> D \<and> M \<subseteq> D})"
@@ -2309,6 +1555,11 @@
     using assms by (auto simp: dynkin_def)
 qed
 
+subsubsection {* Induction rule for intersection-stable generators *}
+
+text {* The reason to introduce Dynkin-systems is the following induction rules for $\sigma$-algebras
+generated by a generator closed under intersection. *}
+
 lemma sigma_sets_induct_disjoint[consumes 3, case_names basic empty compl union]:
   assumes "Int_stable G"
     and closed: "G \<subseteq> Pow \<Omega>"
@@ -2330,4 +1581,756 @@
   with A show ?thesis by auto
 qed
 
+subsection {* 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 'a measure = "{(\<Omega>::'a set, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu> }"
+proof
+  have "sigma_algebra UNIV {{}, UNIV}"
+    by (auto simp: sigma_algebra_iff2)
+  then show "(UNIV, {{}, UNIV}, \<lambda>A. 0) \<in> {(\<Omega>, A, \<mu>). (\<forall>a\<in>-A. \<mu> a = 0) \<and> measure_space \<Omega> A \<mu>} "
+    by (auto simp: measure_space_def positive_def countably_additive_def)
+qed
+
+definition space :: "'a measure \<Rightarrow> 'a set" where
+  "space M = fst (Rep_measure M)"
+
+definition sets :: "'a measure \<Rightarrow> 'a set set" where
+  "sets M = fst (snd (Rep_measure M))"
+
+definition emeasure :: "'a measure \<Rightarrow> 'a set \<Rightarrow> 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 sets!: sigma_algebra "space M" "sets M" for M :: "'a measure"
+  using measure_space[of M] by (auto simp: measure_space_def)
+
+definition measure_of :: "'a set \<Rightarrow> 'a set set \<Rightarrow> ('a set \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
+  "measure_of \<Omega> A \<mu> = Abs_measure (\<Omega>, if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>},
+    \<lambda>a. if a \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> a else 0)"
+
+abbreviation "sigma \<Omega> A \<equiv> measure_of \<Omega> A (\<lambda>x. 0)"
+
+lemma measure_space_0: "A \<subseteq> Pow \<Omega> \<Longrightarrow> measure_space \<Omega> (sigma_sets \<Omega> A) (\<lambda>x. 0)"
+  unfolding measure_space_def
+  by (auto intro!: sigma_algebra_sigma_sets simp: positive_def countably_additive_def)
+
+lemma sigma_algebra_trivial: "sigma_algebra \<Omega> {{}, \<Omega>}"
+by unfold_locales(fastforce intro: exI[where x="{{}}"] exI[where x="{\<Omega>}"])+
+
+lemma measure_space_0': "measure_space \<Omega> {{}, \<Omega>} (\<lambda>x. 0)"
+by(simp add: measure_space_def positive_def countably_additive_def sigma_algebra_trivial)
+
+lemma measure_space_closed:
+  assumes "measure_space \<Omega> M \<mu>"
+  shows "M \<subseteq> Pow \<Omega>"
+proof -
+  interpret sigma_algebra \<Omega> M using assms by(simp add: measure_space_def)
+  show ?thesis by(rule space_closed)
+qed
+
+lemma (in ring_of_sets) positive_cong_eq:
+  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> positive M \<mu>' = positive M \<mu>"
+  by (auto simp add: positive_def)
+
+lemma (in sigma_algebra) countably_additive_eq:
+  "(\<And>a. a \<in> M \<Longrightarrow> \<mu>' a = \<mu> a) \<Longrightarrow> countably_additive M \<mu>' = countably_additive M \<mu>"
+  unfolding countably_additive_def
+  by (intro arg_cong[where f=All] ext) (auto simp add: countably_additive_def subset_eq)
+
+lemma measure_space_eq:
+  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a"
+  shows "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
+proof -
+  interpret sigma_algebra \<Omega> "sigma_sets \<Omega> A" using closed by (rule sigma_algebra_sigma_sets)
+  from positive_cong_eq[OF eq, of "\<lambda>i. i"] countably_additive_eq[OF eq, of "\<lambda>i. i"] show ?thesis
+    by (auto simp: measure_space_def)
+qed
+
+lemma measure_of_eq:
+  assumes closed: "A \<subseteq> Pow \<Omega>" and eq: "(\<And>a. a \<in> sigma_sets \<Omega> A \<Longrightarrow> \<mu> a = \<mu>' a)"
+  shows "measure_of \<Omega> A \<mu> = measure_of \<Omega> A \<mu>'"
+proof -
+  have "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>'"
+    using assms by (rule measure_space_eq)
+  with eq show ?thesis
+    by (auto simp add: measure_of_def intro!: arg_cong[where f=Abs_measure])
+qed
+
+lemma
+  shows space_measure_of_conv: "space (measure_of \<Omega> A \<mu>) = \<Omega>" (is ?space)
+  and sets_measure_of_conv:
+  "sets (measure_of \<Omega> A \<mu>) = (if A \<subseteq> Pow \<Omega> then sigma_sets \<Omega> A else {{}, \<Omega>})" (is ?sets)
+  and emeasure_measure_of_conv: 
+  "emeasure (measure_of \<Omega> A \<mu>) = 
+  (\<lambda>B. if B \<in> sigma_sets \<Omega> A \<and> measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> then \<mu> B else 0)" (is ?emeasure)
+proof -
+  have "?space \<and> ?sets \<and> ?emeasure"
+  proof(cases "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu>")
+    case True
+    from measure_space_closed[OF this] sigma_sets_superset_generator[of A \<Omega>]
+    have "A \<subseteq> Pow \<Omega>" by simp
+    hence "measure_space \<Omega> (sigma_sets \<Omega> A) \<mu> = measure_space \<Omega> (sigma_sets \<Omega> A)
+      (\<lambda>a. if a \<in> sigma_sets \<Omega> A then \<mu> a else 0)"
+      by(rule measure_space_eq) auto
+    with True `A \<subseteq> Pow \<Omega>` show ?thesis
+      by(simp add: measure_of_def space_def sets_def emeasure_def Abs_measure_inverse)
+  next
+    case False thus ?thesis
+      by(cases "A \<subseteq> Pow \<Omega>")(simp_all add: Abs_measure_inverse measure_of_def sets_def space_def emeasure_def measure_space_0 measure_space_0')
+  qed
+  thus ?space ?sets ?emeasure by simp_all
+qed
+
+lemma [simp]:
+  assumes A: "A \<subseteq> Pow \<Omega>"
+  shows sets_measure_of: "sets (measure_of \<Omega> A \<mu>) = sigma_sets \<Omega> A"
+    and space_measure_of: "space (measure_of \<Omega> A \<mu>) = \<Omega>"
+using assms
+by(simp_all add: sets_measure_of_conv space_measure_of_conv)
+
+lemma (in sigma_algebra) sets_measure_of_eq[simp]: "sets (measure_of \<Omega> M \<mu>) = M"
+  using space_closed by (auto intro!: sigma_sets_eq)
+
+lemma (in sigma_algebra) space_measure_of_eq[simp]: "space (measure_of \<Omega> M \<mu>) = \<Omega>"
+  by (rule space_measure_of_conv)
+
+lemma measure_of_subset: "M \<subseteq> Pow \<Omega> \<Longrightarrow> M' \<subseteq> M \<Longrightarrow> sets (measure_of \<Omega> M' \<mu>) \<subseteq> sets (measure_of \<Omega> M \<mu>')"
+  by (auto intro!: sigma_sets_subseteq)
+
+lemma sigma_sets_mono'':
+  assumes "A \<in> sigma_sets C D"
+  assumes "B \<subseteq> D"
+  assumes "D \<subseteq> Pow C"
+  shows "sigma_sets A B \<subseteq> sigma_sets C D"
+proof
+  fix x assume "x \<in> sigma_sets A B"
+  thus "x \<in> sigma_sets C D"
+  proof induct
+    case (Basic a) with assms have "a \<in> D" by auto
+    thus ?case ..
+  next
+    case Empty show ?case by (rule sigma_sets.Empty)
+  next
+    from assms have "A \<in> sets (sigma C D)" by (subst sets_measure_of[OF `D \<subseteq> Pow C`])
+    moreover case (Compl a) hence "a \<in> sets (sigma C D)" by (subst sets_measure_of[OF `D \<subseteq> Pow C`])
+    ultimately have "A - a \<in> sets (sigma C D)" ..
+    thus ?case by (subst (asm) sets_measure_of[OF `D \<subseteq> Pow C`])
+  next
+    case (Union a)
+    thus ?case by (intro sigma_sets.Union)
+  qed
+qed
+
+lemma in_measure_of[intro, simp]: "M \<subseteq> Pow \<Omega> \<Longrightarrow> A \<in> M \<Longrightarrow> A \<in> sets (measure_of \<Omega> M \<mu>)"
+  by auto
+
+subsubsection {* 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)
+  thus ?thesis using X ms
+    by(simp add: M emeasure_measure_of_conv sets_measure_of_conv)
+qed
+
+lemma emeasure_measure_of_sigma:
+  assumes ms: "sigma_algebra \<Omega> M" "positive M \<mu>" "countably_additive M \<mu>"
+  assumes A: "A \<in> M"
+  shows "emeasure (measure_of \<Omega> M \<mu>) A = \<mu> A"
+proof -
+  interpret sigma_algebra \<Omega> M by fact
+  have "measure_space \<Omega> (sigma_sets \<Omega> M) \<mu>"
+    using ms sigma_sets_eq by (simp add: measure_space_def)
+  thus ?thesis by(simp add: emeasure_measure_of_conv A)
+qed
+
+lemma measure_cases[cases type: measure]:
+  obtains (measure) \<Omega> A \<mu> where "x = Abs_measure (\<Omega>, A, \<mu>)" "\<forall>a\<in>-A. \<mu> a = 0" "measure_space \<Omega> A \<mu>"
+  by atomize_elim (cases x, auto)
+
+lemma sets_eq_imp_space_eq:
+  "sets M = sets M' \<Longrightarrow> space M = space M'"
+  using sets.top[of M] sets.top[of M'] sets.space_closed[of M] sets.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 emeasure_neq_0_sets: "emeasure M A \<noteq> 0 \<Longrightarrow> A \<in> sets M"
+  using emeasure_notin_sets[of A M] by blast
+
+lemma measure_notin_sets: "A \<notin> sets M \<Longrightarrow> measure M A = 0"
+  by (simp add: measure_def emeasure_notin_sets)
+
+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 AA': "A = A'" by simp
+  moreover from M.top N.top M.space_closed N.space_closed AA' have "\<Omega> = \<Omega>'" by auto
+  moreover { fix B have "\<mu> B = \<mu>' B"
+    proof cases
+      assume "B \<in> A"
+      with eq `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)
+
+subsubsection {* 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_sets_Collect:
+  assumes f: "f \<in> measurable M N" and P: "{x\<in>space N. P x} \<in> sets N" shows "{x\<in>space M. P (f x)} \<in> sets M"
+proof -
+  have "f -` {x \<in> space N. P x} \<inter> space M = {x\<in>space M. P (f x)}"
+    using measurable_space[OF f] by auto
+  with measurable_sets[OF f P] show ?thesis
+    by simp
+qed
+
+lemma measurable_sigma_sets:
+  assumes B: "sets N = sigma_sets \<Omega> A" "A \<subseteq> Pow \<Omega>"
+      and f: "f \<in> space M \<rightarrow> \<Omega>"
+      and ba: "\<And>y. y \<in> A \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
+  shows "f \<in> measurable M N"
+proof -
+  interpret A: sigma_algebra \<Omega> "sigma_sets \<Omega> A" using B(2) by (rule sigma_algebra_sigma_sets)
+  from B sets.top[of N] A.top sets.space_closed[of N] A.space_closed have \<Omega>: "\<Omega> = space N" by force
+  
+  { fix X assume "X \<in> sigma_sets \<Omega> A"
+    then have "f -` X \<inter> space M \<in> sets M \<and> X \<subseteq> \<Omega>"
+      proof induct
+        case (Basic a) then show ?case
+          by (auto simp add: ba) (metis B(2) subsetD PowD)
+      next
+        case (Compl a)
+        have [simp]: "f -` \<Omega> \<inter> space M = space M"
+          by (auto simp add: funcset_mem [OF f])
+        then show ?case
+          by (auto simp add: vimage_Diff Diff_Int_distrib2 sets.compl_sets Compl)
+      next
+        case (Union a)
+        then show ?case
+          by (simp add: vimage_UN, simp only: UN_extend_simps(4)) blast
+      qed auto }
+  with f show ?thesis
+    by (auto simp add: measurable_def B \<Omega>)
+qed
+
+lemma measurable_measure_of:
+  assumes B: "N \<subseteq> Pow \<Omega>"
+      and f: "f \<in> space M \<rightarrow> \<Omega>"
+      and ba: "\<And>y. y \<in> N \<Longrightarrow> (f -` y) \<inter> space M \<in> sets M"
+  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>)"
+proof -
+  have "sets (measure_of \<Omega> N \<mu>) = sigma_sets \<Omega> N"
+    using B by (rule sets_measure_of)
+  from this assms show ?thesis by (rule measurable_sigma_sets)
+qed
+
+lemma measurable_iff_measure_of:
+  assumes "N \<subseteq> Pow \<Omega>" "f \<in> space M \<rightarrow> \<Omega>"
+  shows "f \<in> measurable M (measure_of \<Omega> N \<mu>) \<longleftrightarrow> (\<forall>A\<in>N. f -` A \<inter> space M \<in> sets M)"
+  by (metis assms in_measure_of measurable_measure_of assms measurable_sets)
+
+lemma measurable_cong_sets:
+  assumes sets: "sets M = sets M'" "sets N = sets N'"
+  shows "measurable M N = measurable M' N'"
+  using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def)
+
+lemma measurable_cong:
+  assumes "\<And> w. w \<in> space M \<Longrightarrow> f w = g w"
+  shows "f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable M M'"
+  unfolding measurable_def using assms
+  by (simp cong: vimage_inter_cong Pi_cong)
+
+lemma measurable_cong_strong:
+  "M = N \<Longrightarrow> M' = N' \<Longrightarrow> (\<And>w. w \<in> space M \<Longrightarrow> f w = g w) \<Longrightarrow>
+    f \<in> measurable M M' \<longleftrightarrow> g \<in> measurable N N'"
+  by (metis measurable_cong)
+
+lemma measurable_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_compose:
+  assumes f: "f \<in> measurable M N" and g: "g \<in> measurable N L"
+  shows "(\<lambda>x. g (f x)) \<in> measurable M L"
+proof -
+  have "\<And>A. (\<lambda>x. g (f x)) -` A \<inter> space M = f -` (g -` A \<inter> space N) \<inter> space M"
+    using measurable_space[OF f] by auto
+  with measurable_space[OF f] measurable_space[OF g] show ?thesis
+    by (auto intro: measurable_sets[OF f] measurable_sets[OF g]
+             simp del: vimage_Int simp add: measurable_def)
+qed
+
+lemma measurable_comp:
+  "f \<in> measurable M N \<Longrightarrow> g \<in> measurable N L \<Longrightarrow> g \<circ> f \<in> measurable M L"
+  using measurable_compose[of f M N g L] by (simp add: comp_def)
+
+lemma measurable_const:
+  "c \<in> space M' \<Longrightarrow> (\<lambda>x. c) \<in> measurable M M'"
+  by (auto simp add: measurable_def)
+
+lemma measurable_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!: sets.Un)
+qed
+
+lemma measurable_If_set:
+  assumes measure: "f \<in> measurable M M'" "g \<in> measurable M M'"
+  assumes P: "A \<inter> space M \<in> sets M"
+  shows "(\<lambda>x. if x \<in> A then f x else g x) \<in> measurable M M'"
+proof (rule measurable_If[OF measure])
+  have "{x \<in> space M. x \<in> A} = A \<inter> space M" by auto
+  thus "{x \<in> space M. x \<in> A} \<in> sets M" using `A \<inter> space M \<in> sets M` by auto
+qed
+
+lemma measurable_ident: "id \<in> measurable M M"
+  by (auto simp add: measurable_def)
+
+lemma measurable_ident_sets:
+  assumes eq: "sets M = sets M'" shows "(\<lambda>x. x) \<in> measurable M M'"
+  using measurable_ident[of M]
+  unfolding id_def measurable_def eq sets_eq_imp_space_eq[OF eq] .
+
+lemma sets_Least:
+  assumes meas: "\<And>i::nat. {x\<in>space M. P i x} \<in> M"
+  shows "(\<lambda>x. LEAST j. P j x) -` A \<inter> space M \<in> sets M"
+proof -
+  { fix i have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M \<in> sets M"
+    proof cases
+      assume i: "(LEAST j. False) = i"
+      have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
+        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x})) \<union> (space M - (\<Union>i. {x\<in>space M. P i x}))"
+        by (simp add: set_eq_iff, safe)
+           (insert i, auto dest: Least_le intro: LeastI intro!: Least_equality)
+      with meas show ?thesis
+        by (auto intro!: sets.Int)
+    next
+      assume i: "(LEAST j. False) \<noteq> i"
+      then have "(\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M =
+        {x\<in>space M. P i x} \<inter> (space M - (\<Union>j<i. {x\<in>space M. P j x}))"
+      proof (simp add: set_eq_iff, safe)
+        fix x assume neq: "(LEAST j. False) \<noteq> (LEAST j. P j x)"
+        have "\<exists>j. P j x"
+          by (rule ccontr) (insert neq, auto)
+        then show "P (LEAST j. P j x) x" by (rule LeastI_ex)
+      qed (auto dest: Least_le intro!: Least_equality)
+      with meas show ?thesis
+        by auto
+    qed }
+  then have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) \<in> sets M"
+    by (intro sets.countable_UN) auto
+  moreover have "(\<Union>i\<in>A. (\<lambda>x. LEAST j. P j x) -` {i} \<inter> space M) =
+    (\<lambda>x. LEAST j. P j x) -` A \<inter> space M" by auto
+  ultimately show ?thesis by auto
+qed
+
+lemma measurable_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. (g \<circ> f) -` 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_comp)
+    apply (metis funcset_mem fab g)
+    apply (subst eq)
+    apply (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)
+
+subsubsection {* 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:
+  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!: sets.finite_UN simp del: UN_simps) }
+  then show ?thesis
+    unfolding measurable_def by auto
+qed
+
+lemma measurable_count_space_eq2_countable:
+  fixes f :: "'a => 'c::countable"
+  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
+proof -
+  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
+    assume *: "\<And>a. a\<in>A \<Longrightarrow> f -` {a} \<inter> space M \<in> sets M"
+    have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)"
+      by auto
+    also have "\<dots> \<in> sets M"
+      using * `X \<subseteq> A` by (intro sets.countable_UN) auto
+    finally have "f -` X \<inter> space M \<in> sets M" . }
+  then show ?thesis
+    unfolding measurable_def by auto
+qed
+
+lemma measurable_compose_countable:
+  assumes f: "\<And>i::'i::countable. (\<lambda>x. f i x) \<in> measurable M N" and g: "g \<in> measurable M (count_space UNIV)"
+  shows "(\<lambda>x. f (g x) x) \<in> measurable M N"
+  unfolding measurable_def
+proof safe
+  fix x assume "x \<in> space M" then show "f (g x) x \<in> space N"
+    using f[THEN measurable_space] g[THEN measurable_space] by auto
+next
+  fix A assume A: "A \<in> sets N"
+  have "(\<lambda>x. f (g x) x) -` A \<inter> space M = (\<Union>i. (g -` {i} \<inter> space M) \<inter> (f i -` A \<inter> space M))"
+    by auto
+  also have "\<dots> \<in> sets M" using f[THEN measurable_sets, OF A] g[THEN measurable_sets]
+    by (auto intro!: sets.countable_UN measurable_sets)
+  finally show "(\<lambda>x. f (g x) x) -` A \<inter> space M \<in> sets M" .
+qed
+
+lemma measurable_count_space_const:
+  "(\<lambda>x. c) \<in> measurable M (count_space UNIV)"
+  by (simp add: measurable_const)
+
+lemma measurable_count_space:
+  "f \<in> measurable (count_space A) (count_space UNIV)"
+  by simp
+
+lemma measurable_compose_rev:
+  assumes f: "f \<in> measurable L N" and g: "g \<in> measurable M L"
+  shows "(\<lambda>x. f (g x)) \<in> measurable M N"
+  using measurable_compose[OF g f] .
+
+lemma measurable_count_space_eq_countable:
+  assumes "countable A"
+  shows "f \<in> measurable M (count_space A) \<longleftrightarrow> (f \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M))"
+proof -
+  { fix X assume "X \<subseteq> A" "f \<in> space M \<rightarrow> A"
+    with `countable A` have "f -` X \<inter> space M = (\<Union>a\<in>X. f -` {a} \<inter> space M)" "countable X"
+      by (auto dest: countable_subset)
+    moreover assume "\<forall>a\<in>A. f -` {a} \<inter> space M \<in> sets M"
+    ultimately have "f -` X \<inter> space M \<in> sets M"
+      using `X \<subseteq> A` by (auto intro!: sets.countable_UN' simp del: UN_simps) }
+  then show ?thesis
+    unfolding measurable_def by auto
+qed
+
+subsubsection {* 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)
+
+subsubsection {* Sigma algebra generated by function preimages *}
+
+definition
+  "vimage_algebra M S X = sigma S ((\<lambda>A. X -` A \<inter> S) ` sets M)"
+
+lemma sigma_algebra_preimages:
+  fixes f :: "'x \<Rightarrow> 'a"
+  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 A assume "A \<in> sets M"
+  moreover have "S - ?F A = ?F (space M - A)"
+    using assms by auto
+  ultimately show "S - ?F A \<in> ?F ` sets M"
+    by blast
+next
+  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 "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> M" "\<And>i. A i = ?F (b i)"
+    by auto
+  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 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 "f \<in> measurable (vimage_algebra M S f) M"
+  unfolding measurable_def using assms by force
+
+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 M S f) M2"
+proof -
+  note measurable_vimage_algebra[OF assms(2)]
+  from measurable_comp[OF this assms(1)]
+  show ?thesis by (simp add: comp_def)
+qed
+
+lemma sigma_sets_vimage:
+  assumes "f \<in> S' \<rightarrow> S" and "A \<subseteq> Pow S"
+  shows "sigma_sets S' ((\<lambda>X. f -` X \<inter> S') ` A) = (\<lambda>X. f -` X \<inter> S') ` sigma_sets S A"
+proof (intro set_eqI iffI)
+  let ?F = "\<lambda>X. f -` X \<inter> S'"
+  fix X assume "X \<in> sigma_sets S' (?F ` A)"
+  then show "X \<in> ?F ` sigma_sets S A"
+  proof induct
+    case (Basic X) then obtain X' where "X = ?F X'" "X' \<in> A"
+      by auto
+    then show ?case by auto
+  next
+    case Empty then show ?case
+      by (auto intro!: image_eqI[of _ _ "{}"] sigma_sets.Empty)
+  next
+    case (Compl X) then obtain X' where X: "X = ?F X'" and "X' \<in> sigma_sets S A"
+      by auto
+    then have "S - X' \<in> sigma_sets S A"
+      by (auto intro!: sigma_sets.Compl)
+    then show ?case
+      using X assms by (auto intro!: image_eqI[where x="S - X'"])
+  next
+    case (Union F)
+    then have "\<forall>i. \<exists>F'.  F' \<in> sigma_sets S A \<and> F i = f -` F' \<inter> S'"
+      by (auto simp: image_iff Bex_def)
+    from choice[OF this] obtain F' where
+      "\<And>i. F' i \<in> sigma_sets S A" and "\<And>i. F i = f -` F' i \<inter> S'"
+      by auto
+    then show ?case
+      by (auto intro!: sigma_sets.Union image_eqI[where x="\<Union>i. F' i"])
+  qed
+next
+  let ?F = "\<lambda>X. f -` X \<inter> S'"
+  fix X assume "X \<in> ?F ` sigma_sets S A"
+  then obtain X' where "X' \<in> sigma_sets S A" "X = ?F X'" by auto
+  then show "X \<in> sigma_sets S' (?F ` A)"
+  proof (induct arbitrary: X)
+    case Empty then show ?case by (auto intro: sigma_sets.Empty)
+  next
+    case (Compl X')
+    have "S' - (S' - X) \<in> sigma_sets S' (?F ` A)"
+      apply (rule sigma_sets.Compl)
+      using assms by (auto intro!: Compl.hyps simp: Compl.prems)
+    also have "S' - (S' - X) = X"
+      using assms Compl by auto
+    finally show ?case .
+  next
+    case (Union F)
+    have "(\<Union>i. f -` F i \<inter> S') \<in> sigma_sets S' (?F ` A)"
+      by (intro sigma_sets.Union Union.hyps) simp
+    also have "(\<Union>i. f -` F i \<inter> S') = X"
+      using assms Union by auto
+    finally show ?case .
+  qed auto
+qed
+
+subsubsection {* Restricted Space Sigma Algebra *}
+
+definition "restrict_space M \<Omega> = measure_of \<Omega> ((op \<inter> \<Omega>) ` sets M) (\<lambda>A. emeasure M (A \<inter> \<Omega>))"
+
+lemma space_restrict_space: "space (restrict_space M \<Omega>) = \<Omega>"
+  unfolding restrict_space_def by (subst space_measure_of) auto
+
+lemma sets_restrict_space: "\<Omega> \<subseteq> space M \<Longrightarrow> sets (restrict_space M \<Omega>) = (op \<inter> \<Omega>) ` sets M"
+  using sigma_sets_vimage[of "\<lambda>x. x" \<Omega> "space M" "sets M"]
+  unfolding restrict_space_def
+  by (subst sets_measure_of) (auto simp: sets.sigma_sets_eq Int_commute[of _ \<Omega>] sets.space_closed)
+
+lemma sets_restrict_space_iff:
+  "\<Omega> \<in> sets M \<Longrightarrow> A \<in> sets (restrict_space M \<Omega>) \<longleftrightarrow> (A \<subseteq> \<Omega> \<and> A \<in> sets M)"
+  by (subst sets_restrict_space) (auto dest: sets.sets_into_space)
+
+lemma measurable_restrict_space1:
+  assumes \<Omega>: "\<Omega> \<in> sets M" and f: "f \<in> measurable M N" shows "f \<in> measurable (restrict_space M \<Omega>) N"
+  unfolding measurable_def
+proof (intro CollectI conjI ballI)
+  show sp: "f \<in> space (restrict_space M \<Omega>) \<rightarrow> space N"
+    using measurable_space[OF f] sets.sets_into_space[OF \<Omega>] by (auto simp: space_restrict_space)
+
+  fix A assume "A \<in> sets N"
+  have "f -` A \<inter> space (restrict_space M \<Omega>) = (f -` A \<inter> space M) \<inter> \<Omega>"
+    using sets.sets_into_space[OF \<Omega>] by (auto simp: space_restrict_space)
+  also have "\<dots> \<in> sets (restrict_space M \<Omega>)"
+    unfolding sets_restrict_space_iff[OF \<Omega>]
+    using measurable_sets[OF f `A \<in> sets N`] \<Omega> by blast
+  finally show "f -` A \<inter> space (restrict_space M \<Omega>) \<in> sets (restrict_space M \<Omega>)" .
+qed
+
+lemma measurable_restrict_space2:
+  "\<Omega> \<in> sets N \<Longrightarrow> f \<in> space M \<rightarrow> \<Omega> \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> f \<in> measurable M (restrict_space N \<Omega>)"
+  by (simp add: measurable_def space_restrict_space sets_restrict_space_iff)
+
 end
--- a/src/HOL/Probability/document/root.tex	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/document/root.tex	Mon May 19 13:44:13 2014 +0200
@@ -13,7 +13,7 @@
 
 \begin{document}
 
-\title{Multivariate Analysis}
+\title{Measure and Probability Theory}
 \maketitle
 
 \tableofcontents
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy	Mon May 19 13:44:13 2014 +0200
@@ -8,7 +8,7 @@
 lemma Ex1_eq: "\<exists>!x. P x \<Longrightarrow> P x \<Longrightarrow> P y \<Longrightarrow> x = y"
   by auto
 
-section "Define the state space"
+subsection {* Define the state space *}
 
 text {*
 
--- a/src/HOL/ROOT	Mon May 19 12:04:45 2014 +0200
+++ b/src/HOL/ROOT	Mon May 19 13:44:13 2014 +0200
@@ -685,8 +685,9 @@
   options [document_graph]
   theories [document = false]
     "~~/src/HOL/Library/Countable"
-    "~~/src/HOL/Multivariate_Analysis/Extended_Real_Limits"
     "~~/src/HOL/Library/Permutation"
+    "~~/src/HOL/Library/Order_Continuity"
+    "~~/src/HOL/Library/Diagonal_Subsequence"
   theories
     Probability
     "ex/Dining_Cryptographers"