merged
authorhoelzl
Thu, 02 Sep 2010 20:44:33 +0200
changeset 39100 e9467adb8b52
parent 39079 bddc3d3f6e53 (current diff)
parent 39099 5c714fd55b1e (diff)
child 39101 606432dd1896
merged
--- a/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Multivariate_Analysis/Topology_Euclidean_Space.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -5027,7 +5027,7 @@
   (\<forall>a\<in>s. \<forall>b\<in>s. \<forall>x. (\<forall>i<DIM('a). ((a$$i \<le> x$$i \<and> x$$i \<le> b$$i) \<or> (b$$i \<le> x$$i \<and> x$$i \<le> a$$i))) \<longrightarrow> x \<in> s)"
 
 lemma is_interval_interval: "is_interval {a .. b::'a::ordered_euclidean_space}" (is ?th1)
-  "is_interval {a<..<b}" (is ?th2) proof - 
+  "is_interval {a<..<b}" (is ?th2) proof -
   have *:"\<And>x y z::real. x < y \<Longrightarrow> y < z \<Longrightarrow> x < z" by auto
   show ?th1 ?th2  unfolding is_interval_def mem_interval Ball_def atLeastAtMost_iff
     by(meson order_trans le_less_trans less_le_trans *)+ qed
@@ -5051,6 +5051,9 @@
 lemma continuous_at_inner: "continuous (at x) (inner a)"
   unfolding continuous_at by (intro tendsto_intros)
 
+lemma continuous_at_euclidean_component[intro!, simp]: "continuous (at x) (\<lambda>x. x $$ i)"
+  unfolding euclidean_component_def by (rule continuous_at_inner)
+
 lemma continuous_on_inner:
   fixes s :: "'a::real_inner set"
   shows "continuous_on s (inner a)"
@@ -5159,6 +5162,9 @@
     by (simp add: closed_def open_halfspace_component_lt)
 qed
 
+lemma open_vimage_euclidean_component: "open S \<Longrightarrow> open ((\<lambda>x. x $$ i) -` S)"
+  by (auto intro!: continuous_open_vimage)
+
 text{* This gives a simple derivation of limit component bounds.                 *}
 
 lemma Lim_component_le: fixes f :: "'a \<Rightarrow> 'b::euclidean_space"
--- a/src/HOL/Probability/Borel.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Borel.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -6,6 +6,10 @@
   imports Sigma_Algebra Positive_Infinite_Real Multivariate_Analysis
 begin
 
+lemma LIMSEQ_max:
+  "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
+  by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
+
 section "Generic Borel spaces"
 
 definition "borel_space = sigma (UNIV::'a::topological_space set) open"
@@ -81,7 +85,7 @@
   "(\<lambda>x. c) \<in> borel_measurable M"
   by (auto intro!: measurable_const)
 
-lemma (in sigma_algebra) borel_measurable_indicator:
+lemma (in sigma_algebra) borel_measurable_indicator[simp, intro!]:
   assumes A: "A \<in> sets M"
   shows "indicator A \<in> borel_measurable M"
   unfolding indicator_def_raw using A
@@ -105,6 +109,53 @@
   qed (auto simp add: vimage_UN)
 qed
 
+lemma (in sigma_algebra) borel_measurable_restricted:
+  fixes f :: "'a \<Rightarrow> 'x\<Colon>{topological_space, semiring_1}" assumes "A \<in> sets M"
+  shows "f \<in> borel_measurable (restricted_space A) \<longleftrightarrow>
+    (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
+    (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
+proof -
+  interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
+  have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R"
+    by (auto intro!: measurable_cong)
+  show ?thesis unfolding *
+    unfolding in_borel_measurable_borel_space
+  proof (simp, safe)
+    fix S :: "'x set" assume "S \<in> sets borel_space"
+      "\<forall>S\<in>sets borel_space. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
+    then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
+    then have f: "?f -` S \<inter> A \<in> sets M"
+      using `A \<in> sets M` sets_into_space by fastsimp
+    show "?f -` S \<inter> space M \<in> sets M"
+    proof cases
+      assume "0 \<in> S"
+      then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)"
+        using `A \<in> sets M` sets_into_space by auto
+      then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff)
+    next
+      assume "0 \<notin> S"
+      then have "?f -` S \<inter> space M = ?f -` S \<inter> A"
+        using `A \<in> sets M` sets_into_space
+        by (auto simp: indicator_def split: split_if_asm)
+      then show ?thesis using f by auto
+    qed
+  next
+    fix S :: "'x set" assume "S \<in> sets borel_space"
+      "\<forall>S\<in>sets borel_space. ?f -` S \<inter> space M \<in> sets M"
+    then have f: "?f -` S \<inter> space M \<in> sets M" by auto
+    then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
+      using `A \<in> sets M` sets_into_space
+      apply (simp add: image_iff)
+      apply (rule bexI[OF _ f])
+      by auto
+  qed
+qed
+
+lemma (in sigma_algebra) borel_measurable_subalgebra:
+  assumes "N \<subseteq> sets M" "f \<in> borel_measurable (M\<lparr>sets:=N\<rparr>)"
+  shows "f \<in> borel_measurable M"
+  using assms unfolding measurable_def by auto
+
 section "Borel spaces on euclidean spaces"
 
 lemma lessThan_borel[simp, intro]:
@@ -658,6 +709,30 @@
   "(f::'a \<Rightarrow> real) \<in> borel_measurable M = (\<forall>a. {w \<in> space M. a < f w} \<in> sets M)"
   using borel_measurable_iff_halfspace_greater[where 'c=real] by simp
 
+lemma borel_measureable_euclidean_component:
+  "(\<lambda>x::'a::euclidean_space. x $$ i) \<in> borel_measurable borel_space"
+  unfolding borel_space_def[where 'a=real]
+proof (rule borel_space.measurable_sigma)
+  fix S::"real set" assume "S \<in> open" then have "open S" unfolding mem_def .
+  from open_vimage_euclidean_component[OF this]
+  show "(\<lambda>x. x $$ i) -` S \<inter> space borel_space \<in> sets borel_space"
+    by (auto intro: borel_space_open)
+qed auto
+
+lemma (in sigma_algebra) borel_measureable_euclidean_space:
+  fixes f :: "'a \<Rightarrow> 'c::ordered_euclidean_space"
+  shows "f \<in> borel_measurable M \<longleftrightarrow> (\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M)"
+proof safe
+  fix i assume "f \<in> borel_measurable M"
+  then show "(\<lambda>x. f x $$ i) \<in> borel_measurable M"
+    using measurable_comp[of f _ _ "\<lambda>x. x $$ i", unfolded comp_def]
+    by (auto intro: borel_measureable_euclidean_component)
+next
+  assume f: "\<forall>i<DIM('c). (\<lambda>x. f x $$ i) \<in> borel_measurable M"
+  then show "f \<in> borel_measurable M"
+    unfolding borel_measurable_iff_halfspace_le by auto
+qed
+
 subsection "Borel measurable operators"
 
 lemma (in sigma_algebra) affine_borel_measurable_vector:
@@ -1270,4 +1345,46 @@
     using assms by auto
 qed
 
+lemma (in sigma_algebra) borel_measurable_psuminf:
+  assumes "\<And>i. f i \<in> borel_measurable M"
+  shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M"
+  using assms unfolding psuminf_def
+  by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand])
+
+section "LIMSEQ is borel measurable"
+
+lemma (in sigma_algebra) borel_measurable_LIMSEQ:
+  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
+  assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
+  and u: "\<And>i. u i \<in> borel_measurable M"
+  shows "u' \<in> borel_measurable M"
+proof -
+  let "?pu x i" = "max (u i x) 0"
+  let "?nu x i" = "max (- u i x) 0"
+
+  { fix x assume x: "x \<in> space M"
+    have "(?pu x) ----> max (u' x) 0"
+      "(?nu x) ----> max (- u' x) 0"
+      using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
+    from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
+    have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
+      "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
+      by (simp_all add: Real_max'[symmetric]) }
+  note eq = this
+
+  have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
+    by auto
+
+  have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M"
+       "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M"
+    using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real)
+  with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space]
+  have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
+       "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M"
+    unfolding SUPR_fun_expand INFI_fun_expand by auto
+  note this[THEN borel_measurable_real]
+  from borel_measurable_diff[OF this]
+  show ?thesis unfolding * .
+qed
+
 end
--- a/src/HOL/Probability/Caratheodory.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Caratheodory.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -445,21 +445,6 @@
     by intro_locales (auto simp add: sigma_algebra_def)
 qed
 
-
-lemma (in algebra) inf_measure_nonempty:
-  assumes f: "positive f" and b: "b \<in> sets M" and a: "a \<subseteq> b"
-  shows "f b \<in> measure_set M f a"
-proof -
-  have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = setsum (f \<circ> (\<lambda>i. {})(0 := b)) {..<1::nat}"
-    by (rule psuminf_finite) (simp add: f[unfolded positive_def])
-  also have "... = f b"
-    by simp
-  finally have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = f b" .
-  thus ?thesis using a b
-    by (auto intro!: exI [of _ "(\<lambda>i. {})(0 := b)"]
-             simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def)
-qed
-
 lemma (in algebra) additive_increasing:
   assumes posf: "positive f" and addf: "additive M f"
   shows "increasing M f"
@@ -494,6 +479,20 @@
     by (auto simp add: Un binaryset_psuminf positive_def)
 qed
 
+lemma inf_measure_nonempty:
+  assumes f: "positive f" and b: "b \<in> sets M" and a: "a \<subseteq> b" "{} \<in> sets M"
+  shows "f b \<in> measure_set M f a"
+proof -
+  have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = setsum (f \<circ> (\<lambda>i. {})(0 := b)) {..<1::nat}"
+    by (rule psuminf_finite) (simp add: f[unfolded positive_def])
+  also have "... = f b"
+    by simp
+  finally have "psuminf (f \<circ> (\<lambda>i. {})(0 := b)) = f b" .
+  thus ?thesis using assms
+    by (auto intro!: exI [of _ "(\<lambda>i. {})(0 := b)"]
+             simp: measure_set_def disjoint_family_on_def split_if_mem2 comp_def)
+qed
+
 lemma (in algebra) inf_measure_agrees:
   assumes posf: "positive f" and ca: "countably_additive M f"
       and s: "s \<in> sets M"
@@ -535,11 +534,11 @@
 qed
 
 lemma (in algebra) inf_measure_empty:
-  assumes posf: "positive f"
+  assumes posf: "positive f"  "{} \<in> sets M"
   shows "Inf (measure_set M f {}) = 0"
 proof (rule antisym)
   show "Inf (measure_set M f {}) \<le> 0"
-    by (metis complete_lattice_class.Inf_lower empty_sets inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def])
+    by (metis complete_lattice_class.Inf_lower `{} \<in> sets M` inf_measure_nonempty[OF posf] subset_refl posf[unfolded positive_def])
 qed simp
 
 lemma (in algebra) inf_measure_positive:
@@ -597,7 +596,7 @@
 next
   case True
   have "measure_set M f s \<noteq> {}"
-    by (metis emptyE ss inf_measure_nonempty [of f, OF posf top])
+    by (metis emptyE ss inf_measure_nonempty [of f, OF posf top _ empty_sets])
   then obtain l where "l \<in> measure_set M f s" by auto
   moreover from True have "l \<le> Inf (measure_set M f s) + e" by simp
   ultimately show ?thesis
--- a/src/HOL/Probability/Information.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Information.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -2,11 +2,53 @@
 imports Probability_Space Product_Measure Convex Radon_Nikodym
 begin
 
+lemma log_le: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x \<le> y \<Longrightarrow> log a x \<le> log a y"
+  by (subst log_le_cancel_iff) auto
+
+lemma log_less: "1 < a \<Longrightarrow> 0 < x \<Longrightarrow> x < y \<Longrightarrow> log a x < log a y"
+  by (subst log_less_cancel_iff) auto
+
+lemma setsum_cartesian_product':
+  "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. setsum (\<lambda>y. f (x, y)) B)"
+  unfolding setsum_cartesian_product by simp
+
 lemma real_of_pinfreal_inverse[simp]:
   fixes X :: pinfreal
   shows "real (inverse X) = 1 / real X"
   by (cases X) (auto simp: inverse_eq_divide)
 
+lemma (in finite_prob_space) finite_product_prob_space_of_images:
+  "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M)\<rparr>
+                     (joint_distribution X Y)"
+  (is "finite_prob_space ?S _")
+proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images)
+  have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
+  thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
+    by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
+qed
+
+lemma (in finite_prob_space) finite_measure_space_prod:
+  assumes X: "finite_measure_space MX (distribution X)"
+  assumes Y: "finite_measure_space MY (distribution Y)"
+  shows "finite_measure_space (prod_measure_space MX MY) (joint_distribution X Y)"
+    (is "finite_measure_space ?M ?D")
+proof (intro finite_measure_spaceI)
+  interpret X: finite_measure_space MX "distribution X" by fact
+  interpret Y: finite_measure_space MY "distribution Y" by fact
+  note finite_measure_space.finite_prod_measure_space[OF X Y, simp]
+  show "finite (space ?M)" using X.finite_space Y.finite_space by auto
+  show "joint_distribution X Y {} = 0" by simp
+  show "sets ?M = Pow (space ?M)" by simp
+  { fix x show "?D (space ?M) \<noteq> \<omega>" by (rule distribution_finite) }
+  { fix A B assume "A \<subseteq> space ?M" "B \<subseteq> space ?M" "A \<inter> B = {}"
+    have *: "(\<lambda>t. (X t, Y t)) -` (A \<union> B) \<inter> space M =
+             (\<lambda>t. (X t, Y t)) -` A \<inter> space M \<union> (\<lambda>t. (X t, Y t)) -` B \<inter> space M"
+      by auto
+    show "?D (A \<union> B) = ?D A + ?D B" unfolding distribution_def *
+      apply (rule measure_additive[symmetric])
+      using `A \<inter> B = {}` by (auto simp: sets_eq_Pow) }
+qed
+
 section "Convex theory"
 
 lemma log_setsum:
@@ -105,51 +147,19 @@
   finally show ?thesis .
 qed
 
-lemma (in finite_prob_space) sum_over_space_distrib:
-  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
-  unfolding distribution_def measure_space_1[symmetric] using finite_space
-  by (subst measure_finitely_additive'')
-     (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\<mu>])
-
-lemma (in finite_prob_space) sum_over_space_real_distribution:
-  "(\<Sum>x\<in>X`space M. real (distribution X {x})) = 1"
-  unfolding distribution_def prob_space[symmetric] using finite_space
-  by (subst real_finite_measure_finite_Union[symmetric])
-     (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob])
+lemma split_pairs:
+  shows
+    "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and
+    "(X = (A, B)) \<longleftrightarrow> (fst X = A \<and> snd X = B)" by auto
 
 section "Information theory"
 
-definition
-  "KL_divergence b M \<mu> \<nu> =
-    measure_space.integral M \<mu> (\<lambda>x. log b (real (sigma_finite_measure.RN_deriv M \<nu> \<mu> x)))"
-
 locale finite_information_space = finite_prob_space +
   fixes b :: real assumes b_gt_1: "1 < b"
 
-lemma (in finite_prob_space) distribution_mono:
-  assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
-  shows "distribution X x \<le> distribution Y y"
-  unfolding distribution_def
-  using assms by (auto simp: sets_eq_Pow intro!: measure_mono)
-
-lemma (in prob_space) distribution_remove_const:
-  shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
-  and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
-  and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
-  and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
-  and "distribution (\<lambda>x. ()) {()} = 1"
-  unfolding measure_space_1[symmetric]
-  by (auto intro!: arg_cong[where f="\<mu>"] simp: distribution_def)
-
 context finite_information_space
 begin
 
-lemma distribution_mono_gt_0:
-  assumes gt_0: "0 < distribution X x"
-  assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
-  shows "0 < distribution Y y"
-  by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
-
 lemma
   assumes "0 \<le> A" and pos: "0 < A \<Longrightarrow> 0 < B" "0 < A \<Longrightarrow> 0 < C"
   shows mult_log_mult: "A * log b (B * C) = A * log b B + A * log b C" (is "?mult")
@@ -165,41 +175,6 @@
   thus ?mult and ?div by auto
 qed
 
-lemma split_pairs:
-  shows
-    "((A, B) = X) \<longleftrightarrow> (fst X = A \<and> snd X = B)" and
-    "(X = (A, B)) \<longleftrightarrow> (fst X = A \<and> snd X = B)" by auto
-
-lemma (in finite_information_space) distribution_finite:
-  "distribution X A \<noteq> \<omega>"
-  using measure_finite[of "X -` A \<inter> space M"]
-  unfolding distribution_def sets_eq_Pow by auto
-
-lemma (in finite_information_space) real_distribution_gt_0[simp]:
-  "0 < real (distribution Y y) \<longleftrightarrow>  0 < distribution Y y"
-  using assms by (auto intro!: real_pinfreal_pos distribution_finite)
-
-lemma real_distribution_mult_pos_pos:
-  assumes "0 < distribution Y y"
-  and "0 < distribution X x"
-  shows "0 < real (distribution Y y * distribution X x)"
-  unfolding real_of_pinfreal_mult[symmetric]
-  using assms by (auto intro!: mult_pos_pos)
-
-lemma real_distribution_divide_pos_pos:
-  assumes "0 < distribution Y y"
-  and "0 < distribution X x"
-  shows "0 < real (distribution Y y / distribution X x)"
-  unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric]
-  using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
-
-lemma real_distribution_mult_inverse_pos_pos:
-  assumes "0 < distribution Y y"
-  and "0 < distribution X x"
-  shows "0 < real (distribution Y y * inverse (distribution X x))"
-  unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric]
-  using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
-
 ML {*
 
   (* tactic to solve equations of the form @{term "W * log b (X / (Y * Z)) = W * log b X - W * log b (Y * Z)"}
@@ -252,31 +227,14 @@
 
 end
 
-lemma (in finite_measure_space) absolutely_continuousI:
-  assumes "finite_measure_space M \<nu>"
-  assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
-  shows "absolutely_continuous \<nu>"
-proof (unfold absolutely_continuous_def sets_eq_Pow, safe)
-  fix N assume "\<mu> N = 0" "N \<subseteq> space M"
-
-  interpret v: finite_measure_space M \<nu> by fact
+subsection "Kullback$-$Leibler divergence"
 
-  have "\<nu> N = \<nu> (\<Union>x\<in>N. {x})" by simp
-  also have "\<dots> = (\<Sum>x\<in>N. \<nu> {x})"
-  proof (rule v.measure_finitely_additive''[symmetric])
-    show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset)
-    show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto
-    fix x assume "x \<in> N" thus "{x} \<in> sets M" using `N \<subseteq> space M` sets_eq_Pow by auto
-  qed
-  also have "\<dots> = 0"
-  proof (safe intro!: setsum_0')
-    fix x assume "x \<in> N"
-    hence "\<mu> {x} \<le> \<mu> N" using sets_eq_Pow `N \<subseteq> space M` by (auto intro!: measure_mono)
-    hence "\<mu> {x} = 0" using `\<mu> N = 0` by simp
-    thus "\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
-  qed
-  finally show "\<nu> N = 0" .
-qed
+text {* The Kullback$-$Leibler divergence is also known as relative entropy or
+Kullback$-$Leibler distance. *}
+
+definition
+  "KL_divergence b M \<mu> \<nu> =
+    measure_space.integral M \<mu> (\<lambda>x. log b (real (sigma_finite_measure.RN_deriv M \<nu> \<mu> x)))"
 
 lemma (in finite_measure_space) KL_divergence_eq_finite:
   assumes v: "finite_measure_space M \<nu>"
@@ -285,19 +243,13 @@
 proof (simp add: KL_divergence_def finite_measure_space.integral_finite_singleton[OF v])
   interpret v: finite_measure_space M \<nu> by fact
   have ms: "measure_space M \<nu>" by fact
-
   have ac: "absolutely_continuous \<nu>"
     using ac by (auto intro!: absolutely_continuousI[OF v])
-
   show "(\<Sum>x \<in> space M. log b (real (RN_deriv \<nu> x)) * real (\<nu> {x})) = ?sum"
     using RN_deriv_finite_measure[OF ms ac]
     by (auto intro!: setsum_cong simp: field_simps real_of_pinfreal_mult[symmetric])
 qed
 
-lemma (in finite_prob_space) finite_sum_over_space_eq_1:
-  "(\<Sum>x\<in>space M. real (\<mu> {x})) = 1"
-  using sum_over_space_eq_1 finite_measure by (simp add: real_of_pinfreal_setsum sets_eq_Pow)
-
 lemma (in finite_prob_space) KL_divergence_positive_finite:
   assumes v: "finite_prob_space M \<nu>"
   assumes ac: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
@@ -322,13 +274,15 @@
       fix x assume x: "x \<in> space M"
       { assume "0 < real (\<nu> {x})"
         hence "\<mu> {x} \<noteq> 0" using ac[OF x] by auto
-        thus "0 < prob {x}" using measure_finite[of "{x}"] sets_eq_Pow x
+        thus "0 < prob {x}" using finite_measure[of "{x}"] sets_eq_Pow x
           by (cases "\<mu> {x}") simp_all }
     qed auto
   qed
   thus "0 \<le> KL_divergence b M \<nu> \<mu>" using finite_sum_over_space_eq_1 by simp
 qed
 
+subsection {* Mutual Information *}
+
 definition (in prob_space)
   "mutual_information b S T X Y =
     KL_divergence b (prod_measure_space S T)
@@ -341,24 +295,48 @@
     \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr>
     \<lparr> space = Y`space M, sets = Pow (Y`space M) \<rparr> X Y"
 
-lemma prod_measure_times_finite:
-  assumes fms: "finite_measure_space M \<mu>" "finite_measure_space N \<nu>" and a: "a \<in> space M \<times> space N"
-  shows "prod_measure M \<mu> N \<nu> {a} = \<mu> {fst a} * \<nu> {snd a}"
-proof (cases a)
-  case (Pair b c)
-  hence a_eq: "{a} = {b} \<times> {c}" by simp
-
-  interpret M: finite_measure_space M \<mu> by fact
-  interpret N: finite_measure_space N \<nu> by fact
-
-  from finite_measure_space.finite_prod_measure_times[OF fms, of "{b}" "{c}"] M.sets_eq_Pow N.sets_eq_Pow a Pair
-  show ?thesis unfolding a_eq by simp
+lemma (in finite_information_space) mutual_information_generic_eq:
+  assumes MX: "finite_measure_space MX (distribution X)"
+  assumes MY: "finite_measure_space MY (distribution Y)"
+  shows "mutual_information b MX MY X Y = (\<Sum> (x,y) \<in> space MX \<times> space MY.
+      real (joint_distribution X Y {(x,y)}) *
+      log b (real (joint_distribution X Y {(x,y)}) /
+      (real (distribution X {x}) * real (distribution Y {y}))))"
+proof -
+  let ?P = "prod_measure_space MX MY"
+  let ?\<mu> = "prod_measure MX (distribution X) MY (distribution Y)"
+  let ?\<nu> = "joint_distribution X Y"
+  interpret X: finite_measure_space MX "distribution X" by fact
+  moreover interpret Y: finite_measure_space MY "distribution Y" by fact
+  have fms: "finite_measure_space MX (distribution X)"
+            "finite_measure_space MY (distribution Y)" by fact+
+  have fms_P: "finite_measure_space ?P ?\<mu>"
+    by (rule X.finite_measure_space_finite_prod_measure) fact
+  then interpret P: finite_measure_space ?P ?\<mu> .
+  have fms_P': "finite_measure_space ?P ?\<nu>"
+      using finite_product_measure_space[of "space MX" "space MY"]
+        X.finite_space Y.finite_space sigma_prod_sets_finite[OF X.finite_space Y.finite_space]
+        X.sets_eq_Pow Y.sets_eq_Pow
+      by (simp add: prod_measure_space_def sigma_def)
+  then interpret P': finite_measure_space ?P ?\<nu> .
+  { fix x assume "x \<in> space ?P"
+    hence in_MX: "{fst x} \<in> sets MX" "{snd x} \<in> sets MY" using X.sets_eq_Pow Y.sets_eq_Pow
+      by (auto simp: prod_measure_space_def)
+    assume "?\<mu> {x} = 0"
+    with X.finite_prod_measure_times[OF fms(2), of "{fst x}" "{snd x}"] in_MX
+    have "distribution X {fst x} = 0 \<or> distribution Y {snd x} = 0"
+      by (simp add: prod_measure_space_def)
+    hence "joint_distribution X Y {x} = 0"
+      by (cases x) (auto simp: distribution_order) }
+  note measure_0 = this
+  show ?thesis
+    unfolding Let_def mutual_information_def
+    using measure_0 fms_P fms_P' MX MY P.absolutely_continuous_def
+    by (subst P.KL_divergence_eq_finite)
+       (auto simp add: prod_measure_space_def prod_measure_times_finite
+         finite_prob_space_eq setsum_cartesian_product' real_of_pinfreal_mult[symmetric])
 qed
 
-lemma setsum_cartesian_product':
-  "(\<Sum>x\<in>A \<times> B. f x) = (\<Sum>x\<in>A. setsum (\<lambda>y. f (x, y)) B)"
-  unfolding setsum_cartesian_product by simp
-
 lemma (in finite_information_space)
   assumes MX: "finite_prob_space MX (distribution X)"
   assumes MY: "finite_prob_space MY (distribution Y)"
@@ -436,9 +414,26 @@
                                                    (real (distribution X {x}) * real (distribution Y {y}))))"
   by (subst mutual_information_eq_generic) (simp_all add: finite_prob_space_of_images)
 
+lemma (in finite_information_space) mutual_information_cong:
+  assumes X: "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
+  assumes Y: "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
+  shows "\<I>(X ; Y) = \<I>(X' ; Y')"
+proof -
+  have "X ` space M = X' ` space M" using X by (auto intro!: image_eqI)
+  moreover have "Y ` space M = Y' ` space M" using Y by (auto intro!: image_eqI)
+  ultimately show ?thesis
+  unfolding mutual_information_eq
+    using
+      assms[THEN distribution_cong]
+      joint_distribution_cong[OF assms]
+    by (auto intro!: setsum_cong)
+qed
+
 lemma (in finite_information_space) mutual_information_positive: "0 \<le> \<I>(X;Y)"
   by (subst mutual_information_positive_generic) (simp_all add: finite_prob_space_of_images)
 
+subsection {* Entropy *}
+
 definition (in prob_space)
   "entropy b s X = mutual_information b s s X X"
 
@@ -446,32 +441,146 @@
   finite_entropy ("\<H>'(_')") where
   "\<H>(X) \<equiv> entropy b \<lparr> space = X`space M, sets = Pow (X`space M) \<rparr> X"
 
-lemma (in finite_information_space) joint_distribution_remove[simp]:
-    "joint_distribution X X {(x, x)} = distribution X {x}"
-  unfolding distribution_def by (auto intro!: arg_cong[where f="\<mu>"])
+lemma (in finite_information_space) entropy_generic_eq:
+  assumes MX: "finite_measure_space MX (distribution X)"
+  shows "entropy b MX X = -(\<Sum> x \<in> space MX. real (distribution X {x}) * log b (real (distribution X {x})))"
+proof -
+  let "?X x" = "real (distribution X {x})"
+  let "?XX x y" = "real (joint_distribution X X {(x, y)})"
+  interpret MX: finite_measure_space MX "distribution X" by fact
+  { fix x y
+    have "(\<lambda>x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto
+    then have "?XX x y * log b (?XX x y / (?X x * ?X y)) =
+        (if x = y then - ?X y * log b (?X y) else 0)"
+      unfolding distribution_def by (auto simp: mult_log_divide) }
+  note remove_XX = this
+  show ?thesis
+    unfolding entropy_def mutual_information_generic_eq[OF MX MX]
+    unfolding setsum_cartesian_product[symmetric] setsum_negf[symmetric] remove_XX
+    by (auto simp: setsum_cases MX.finite_space)
+qed
 
 lemma (in finite_information_space) entropy_eq:
   "\<H>(X) = -(\<Sum> x \<in> X ` space M. real (distribution X {x}) * log b (real (distribution X {x})))"
-proof -
-  { fix f
-    { fix x y
-      have "(\<lambda>x. (X x, X x)) -` {(x, y)} = (if x = y then X -` {x} else {})" by auto
-        hence "real (distribution (\<lambda>x. (X x, X x))  {(x,y)}) * f x y =
-            (if x = y then real (distribution X {x}) * f x y else 0)"
-        unfolding distribution_def by auto }
-    hence "(\<Sum>(x, y) \<in> X ` space M \<times> X ` space M. real (joint_distribution X X {(x, y)}) * f x y) =
-      (\<Sum>x \<in> X ` space M. real (distribution X {x}) * f x x)"
-      unfolding setsum_cartesian_product' by (simp add: setsum_cases finite_space) }
-  note remove_cartesian_product = this
-
-  show ?thesis
-    unfolding entropy_def mutual_information_eq setsum_negf[symmetric] remove_cartesian_product
-    by (auto intro!: setsum_cong)
-qed
+  by (simp add: finite_measure_space entropy_generic_eq)
 
 lemma (in finite_information_space) entropy_positive: "0 \<le> \<H>(X)"
   unfolding entropy_def using mutual_information_positive .
 
+lemma (in finite_information_space) entropy_certainty_eq_0:
+  assumes "x \<in> X ` space M" and "distribution X {x} = 1"
+  shows "\<H>(X) = 0"
+proof -
+  interpret X: finite_prob_space "\<lparr> space = X ` space M, sets = Pow (X ` space M) \<rparr>" "distribution X"
+    by (rule finite_prob_space_of_images)
+
+  have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}"
+    using X.measure_compl[of "{x}"] assms by auto
+  also have "\<dots> = 0" using X.prob_space assms by auto
+  finally have X0: "distribution X (X ` space M - {x}) = 0" by auto
+
+  { fix y assume asm: "y \<noteq> x" "y \<in> X ` space M"
+    hence "{y} \<subseteq> X ` space M - {x}" by auto
+    from X.measure_mono[OF this] X0 asm
+    have "distribution X {y} = 0" by auto }
+
+  hence fi: "\<And> y. y \<in> X ` space M \<Longrightarrow> real (distribution X {y}) = (if x = y then 1 else 0)"
+    using assms by auto
+
+  have y: "\<And>y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp
+
+  show ?thesis unfolding entropy_eq by (auto simp: y fi)
+qed
+
+lemma (in finite_information_space) entropy_le_card_not_0:
+  "\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))"
+proof -
+  let "?d x" = "distribution X {x}"
+  let "?p x" = "real (?d x)"
+  have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))"
+    by (auto intro!: setsum_cong simp: entropy_eq setsum_negf[symmetric])
+  also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))"
+    apply (rule log_setsum')
+    using not_empty b_gt_1 finite_space sum_over_space_real_distribution
+    by auto
+  also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?d x \<noteq> 0 then 1 else 0)"
+    apply (rule arg_cong[where f="\<lambda>f. log b (\<Sum>x\<in>X`space M. f x)"])
+    using distribution_finite[of X] by (auto simp: expand_fun_eq real_of_pinfreal_eq_0)
+  finally show ?thesis
+    using finite_space by (auto simp: setsum_cases real_eq_of_nat)
+qed
+
+lemma (in finite_information_space) entropy_uniform_max:
+  assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}"
+  shows "\<H>(X) = log b (real (card (X ` space M)))"
+proof -
+  note uniform =
+    finite_prob_space_of_images[of X, THEN finite_prob_space.uniform_prob, simplified]
+
+  have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
+    using finite_space not_empty by auto
+
+  { fix x assume "x \<in> X ` space M"
+    hence "real (distribution X {x}) = 1 / real (card (X ` space M))"
+    proof (rule uniform)
+      fix x y assume "x \<in> X`space M" "y \<in> X`space M"
+      from assms[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp
+    qed }
+  thus ?thesis
+    using not_empty finite_space b_gt_1 card_gt0
+    by (simp add: entropy_eq real_eq_of_nat[symmetric] log_divide)
+qed
+
+lemma (in finite_information_space) entropy_le_card:
+  "\<H>(X) \<le> log b (real (card (X ` space M)))"
+proof cases
+  assume "X ` space M \<inter> {x. distribution X {x} \<noteq> 0} = {}"
+  then have "\<And>x. x\<in>X`space M \<Longrightarrow> distribution X {x} = 0" by auto
+  moreover
+  have "0 < card (X`space M)"
+    using finite_space not_empty unfolding card_gt_0_iff by auto
+  then have "log b 1 \<le> log b (real (card (X`space M)))"
+    using b_gt_1 by (intro log_le) auto
+  ultimately show ?thesis unfolding entropy_eq by simp
+next
+  assume False: "X ` space M \<inter> {x. distribution X {x} \<noteq> 0} \<noteq> {}"
+  have "card (X ` space M \<inter> {x. distribution X {x} \<noteq> 0}) \<le> card (X ` space M)"
+    (is "?A \<le> ?B") using finite_space not_empty by (auto intro!: card_mono)
+  note entropy_le_card_not_0
+  also have "log b (real ?A) \<le> log b (real ?B)"
+    using b_gt_1 False finite_space not_empty `?A \<le> ?B`
+    by (auto intro!: log_le simp: card_gt_0_iff)
+  finally show ?thesis .
+qed
+
+lemma (in finite_information_space) entropy_commute:
+  "\<H>(\<lambda>x. (X x, Y x)) = \<H>(\<lambda>x. (Y x, X x))"
+proof -
+  have *: "(\<lambda>x. (Y x, X x))`space M = (\<lambda>(a,b). (b,a))`(\<lambda>x. (X x, Y x))`space M"
+    by auto
+  have inj: "\<And>X. inj_on (\<lambda>(a,b). (b,a)) X"
+    by (auto intro!: inj_onI)
+  show ?thesis
+    unfolding entropy_eq unfolding * setsum_reindex[OF inj]
+    by (simp add: joint_distribution_commute[of Y X] split_beta)
+qed
+
+lemma (in finite_information_space) entropy_eq_cartesian_sum:
+  "\<H>(\<lambda>x. (X x, Y x)) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
+    real (joint_distribution X Y {(x,y)}) *
+    log b (real (joint_distribution X Y {(x,y)})))"
+proof -
+  { fix x assume "x\<notin>(\<lambda>x. (X x, Y x))`space M"
+    then have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M = {}" by auto
+    then have "joint_distribution X Y {x} = 0"
+      unfolding distribution_def by auto }
+  then show ?thesis using finite_space
+    unfolding entropy_eq neg_equal_iff_equal setsum_cartesian_product
+    by (auto intro!: setsum_mono_zero_cong_left)
+qed
+
+subsection {* Conditional Mutual Information *}
+
 definition (in prob_space)
   "conditional_mutual_information b M1 M2 M3 X Y Z \<equiv>
     mutual_information b M1 (prod_measure_space M2 M3) X (\<lambda>x. (Y x, Z x)) -
@@ -485,87 +594,32 @@
     \<lparr> space = Z`space M, sets = Pow (Z`space M) \<rparr>
     X Y Z"
 
-lemma (in finite_information_space) setsum_distribution_gen:
-  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
-  and "inj_on f (X`space M)"
-  shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
-  unfolding distribution_def assms
-  using finite_space assms
-  by (subst measure_finitely_additive'')
-     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
-      intro!: arg_cong[where f=prob])
-
-lemma (in finite_information_space) setsum_distribution:
-  "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}"
-  "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}"
-  "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}"
-  "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}"
-  "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
-  by (auto intro!: inj_onI setsum_distribution_gen)
-
-lemma (in finite_information_space) setsum_real_distribution_gen:
-  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
-  and "inj_on f (X`space M)"
-  shows "(\<Sum>x \<in> X`space M. real (distribution Y {f x})) = real (distribution Z {c})"
-  unfolding distribution_def assms
-  using finite_space assms
-  by (subst real_finite_measure_finite_Union[symmetric])
-     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
-        intro!: arg_cong[where f=prob])
-
-lemma (in finite_information_space) setsum_real_distribution:
-  "(\<Sum>x \<in> X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})"
-  "(\<Sum>y \<in> Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})"
-  "(\<Sum>x \<in> X`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})"
-  "(\<Sum>y \<in> Y`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})"
-  "(\<Sum>z \<in> Z`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})"
-  by (auto intro!: inj_onI setsum_real_distribution_gen)
+lemma (in finite_information_space) conditional_mutual_information_generic_eq:
+  assumes MX: "finite_measure_space MX (distribution X)"
+  assumes MY: "finite_measure_space MY (distribution Y)"
+  assumes MZ: "finite_measure_space MZ (distribution Z)"
+  shows "conditional_mutual_information b MX MY MZ X Y Z =
+    (\<Sum>(x, y, z)\<in>space MX \<times> space MY \<times> space MZ.
+      real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) *
+      log b (real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) /
+                   (real (distribution X {x}) * real (joint_distribution Y Z {(y, z)})))) -
+    (\<Sum>(x, y)\<in>space MX \<times> space MZ.
+      real (joint_distribution X Z {(x, y)}) *
+      log b (real (joint_distribution X Z {(x, y)}) / (real (distribution X {x}) * real (distribution Z {y}))))"
+  using assms finite_measure_space_prod[OF MY MZ]
+  unfolding conditional_mutual_information_def
+  by (subst (1 2) mutual_information_generic_eq)
+     (simp_all add: setsum_cartesian_product' finite_measure_space.finite_prod_measure_space)
 
-lemma (in finite_information_space) conditional_mutual_information_eq_sum:
-   "\<I>(X ; Y | Z) =
-     (\<Sum>(x, y, z)\<in>X ` space M \<times> (\<lambda>x. (Y x, Z x)) ` space M.
-             real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
-             log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)})/
-        real (distribution (\<lambda>x. (Y x, Z x)) {(y, z)}))) -
-     (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M.
-        real (distribution (\<lambda>x. (X x, Z x)) {(x,z)}) * log b (real (distribution (\<lambda>x. (X x, Z x)) {(x,z)}) / real (distribution Z {z})))"
-  (is "_ = ?rhs")
-proof -
-  have setsum_product:
-    "\<And>f x. (\<Sum>v\<in>(\<lambda>x. (Y x, Z x)) ` space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x,v)}) * f v)
-      = (\<Sum>v\<in>Y ` space M \<times> Z ` space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x,v)}) * f v)"
-  proof (safe intro!: setsum_mono_zero_cong_left imageI)
-    fix x y z f
-    assume *: "(Y y, Z z) \<notin> (\<lambda>x. (Y x, Z x)) ` space M" and "y \<in> space M" "z \<in> space M"
-    hence "(\<lambda>x. (X x, Y x, Z x)) -` {(x, Y y, Z z)} \<inter> space M = {}"
-    proof safe
-      fix x' assume x': "x' \<in> space M" and eq: "Y x' = Y y" "Z x' = Z z"
-      have "(Y y, Z z) \<in> (\<lambda>x. (Y x, Z x)) ` space M" using eq[symmetric] x' by auto
-      thus "x' \<in> {}" using * by auto
-    qed
-    thus "real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, Y y, Z z)}) * f (Y y) (Z z) = 0"
-      unfolding distribution_def by simp
-  qed (simp add: finite_space)
-
-  thus ?thesis
-    unfolding conditional_mutual_information_def Let_def mutual_information_eq
-    by (subst mutual_information_eq_generic)
-       (auto simp: prod_measure_space_def sigma_prod_sets_finite finite_space sigma_def
-        finite_prob_space_of_images finite_product_prob_space_of_images
-        setsum_cartesian_product' setsum_product setsum_subtractf setsum_addf
-        setsum_left_distrib[symmetric] setsum_real_distribution
-      cong: setsum_cong)
-qed
 
 lemma (in finite_information_space) conditional_mutual_information_eq:
   "\<I>(X ; Y | Z) = (\<Sum>(x, y, z) \<in> X ` space M \<times> Y ` space M \<times> Z ` space M.
              real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) *
              log b (real (distribution (\<lambda>x. (X x, Y x, Z x)) {(x, y, z)}) /
     (real (joint_distribution X Z {(x, z)}) * real (joint_distribution Y Z {(y,z)} / distribution Z {z}))))"
-  unfolding conditional_mutual_information_def Let_def mutual_information_eq
-  by (subst mutual_information_eq_generic)
+  by (subst conditional_mutual_information_generic_eq)
      (auto simp add: prod_measure_space_def sigma_prod_sets_finite finite_space
-      finite_prob_space_of_images finite_product_prob_space_of_images sigma_def
+      finite_measure_space finite_product_prob_space_of_images sigma_def
       setsum_cartesian_product' setsum_product setsum_subtractf setsum_addf
       setsum_left_distrib[symmetric] setsum_real_distribution setsum_commute[where A="Y`space M"]
       real_of_pinfreal_mult[symmetric]
@@ -581,22 +635,6 @@
     by (simp add: setsum_cartesian_product' distribution_remove_const)
 qed
 
-lemma (in finite_prob_space) distribution_finite:
-  "distribution X A \<noteq> \<omega>"
-  by (auto simp: sets_eq_Pow distribution_def intro!: measure_finite)
-
-lemma (in finite_prob_space) real_distribution_order:
-  shows "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution X {x})"
-  and "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution Y {y})"
-  and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution X {x})"
-  and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution Y {y})"
-  and "distribution X {x} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
-  and "distribution Y {y} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
-  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
-  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
-  using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"]
-  by auto
-
 lemma (in finite_information_space) conditional_mutual_information_positive:
   "0 \<le> \<I>(X ; Y | Z)"
 proof -
@@ -640,6 +678,8 @@
     by (simp add: real_of_pinfreal_mult[symmetric])
 qed
 
+subsection {* Conditional Entropy *}
+
 definition (in prob_space)
   "conditional_entropy b S T X Y = conditional_mutual_information b S S T X X Y"
 
@@ -652,19 +692,69 @@
 lemma (in finite_information_space) conditional_entropy_positive:
   "0 \<le> \<H>(X | Y)" unfolding conditional_entropy_def using conditional_mutual_information_positive .
 
+lemma (in finite_information_space) conditional_entropy_generic_eq:
+  assumes MX: "finite_measure_space MX (distribution X)"
+  assumes MY: "finite_measure_space MZ (distribution Z)"
+  shows "conditional_entropy b MX MZ X Z =
+     - (\<Sum>(x, z)\<in>space MX \<times> space MZ.
+         real (joint_distribution X Z {(x, z)}) *
+         log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))"
+  unfolding conditional_entropy_def using assms
+  apply (simp add: conditional_mutual_information_generic_eq
+                   setsum_cartesian_product' setsum_commute[of _ "space MZ"]
+                   setsum_negf[symmetric] setsum_subtractf[symmetric])
+proof (safe intro!: setsum_cong, simp)
+  fix z x assume "z \<in> space MZ" "x \<in> space MX"
+  let "?XXZ x'" = "real (joint_distribution X (\<lambda>x. (X x, Z x)) {(x, x', z)})"
+  let "?XZ x'" = "real (joint_distribution X Z {(x', z)})"
+  let "?X" = "real (distribution X {x})"
+  interpret MX: finite_measure_space MX "distribution X" by fact
+  have *: "\<And>A. A = {} \<Longrightarrow> prob A = 0" by simp
+  have XXZ: "\<And>x'. ?XXZ x' = (if x' = x then ?XZ x else 0)"
+    by (auto simp: distribution_def intro!: arg_cong[where f=prob] *)
+  have "(\<Sum>x'\<in>space MX. ?XXZ x' * log b (?XXZ x') - (?XXZ x' * log b ?X + ?XXZ x' * log b (?XZ x'))) =
+    (\<Sum>x'\<in>{x}. ?XZ x' * log b (?XZ x') - (?XZ x' * log b ?X + ?XZ x' * log b (?XZ x')))"
+    using `x \<in> space MX` MX.finite_space
+    by (safe intro!: setsum_mono_zero_cong_right)
+       (auto split: split_if_asm simp: XXZ)
+  then show "(\<Sum>x'\<in>space MX. ?XXZ x' * log b (?XXZ x') - (?XXZ x' * log b ?X + ?XXZ x' * log b (?XZ x'))) +
+      ?XZ x * log b ?X = 0" by simp
+qed
+
 lemma (in finite_information_space) conditional_entropy_eq:
   "\<H>(X | Z) =
      - (\<Sum>(x, z)\<in>X ` space M \<times> Z ` space M.
          real (joint_distribution X Z {(x, z)}) *
          log b (real (joint_distribution X Z {(x, z)}) / real (distribution Z {z})))"
+  by (simp add: finite_measure_space conditional_entropy_generic_eq)
+
+lemma (in finite_information_space) conditional_entropy_eq_ce_with_hypothesis:
+  "\<H>(X | Y) =
+    -(\<Sum>y\<in>Y`space M. real (distribution Y {y}) *
+      (\<Sum>x\<in>X`space M. real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}) *
+              log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {(y)}))))"
+  unfolding conditional_entropy_eq neg_equal_iff_equal
+  apply (simp add: setsum_commute[of _ "Y`space M"] setsum_cartesian_product' setsum_divide_distrib[symmetric])
+  apply (safe intro!: setsum_cong)
+  using real_distribution_order'[of Y _ X _]
+  by (auto simp add: setsum_subtractf[of _ _ "X`space M"])
+
+lemma (in finite_information_space) conditional_entropy_eq_cartesian_sum:
+  "\<H>(X | Y) = -(\<Sum>x\<in>X`space M. \<Sum>y\<in>Y`space M.
+    real (joint_distribution X Y {(x,y)}) *
+    log b (real (joint_distribution X Y {(x,y)}) / real (distribution Y {y})))"
 proof -
-  have *: "\<And>x y z. (\<lambda>x. (X x, X x, Z x)) -` {(x, y, z)} = (if x = y then (\<lambda>x. (X x, Z x)) -` {(x, z)} else {})" by auto
-  show ?thesis
-    unfolding conditional_mutual_information_eq_sum
-      conditional_entropy_def distribution_def *
-    by (auto intro!: setsum_0')
+  { fix x assume "x\<notin>(\<lambda>x. (X x, Y x))`space M"
+    then have "(\<lambda>x. (X x, Y x)) -` {x} \<inter> space M = {}" by auto
+    then have "joint_distribution X Y {x} = 0"
+      unfolding distribution_def by auto }
+  then show ?thesis using finite_space
+    unfolding conditional_entropy_eq neg_equal_iff_equal setsum_cartesian_product
+    by (auto intro!: setsum_mono_zero_cong_left)
 qed
 
+subsection {* Equalities *}
+
 lemma (in finite_information_space) mutual_information_eq_entropy_conditional_entropy:
   "\<I>(X ; Z) = \<H>(X) - \<H>(X | Z)"
   unfolding mutual_information_eq entropy_eq conditional_entropy_eq
@@ -680,109 +770,15 @@
   show ?thesis by auto
 qed
 
-(* -------------Entropy of a RV with a certain event is zero---------------- *)
-
-lemma (in finite_information_space) finite_entropy_certainty_eq_0:
-  assumes "x \<in> X ` space M" and "distribution X {x} = 1"
-  shows "\<H>(X) = 0"
-proof -
-  interpret X: finite_prob_space "\<lparr> space = X ` space M, sets = Pow (X ` space M) \<rparr>" "distribution X"
-    by (rule finite_prob_space_of_images)
-
-  have "distribution X (X ` space M - {x}) = distribution X (X ` space M) - distribution X {x}"
-    using X.measure_compl[of "{x}"] assms by auto
-  also have "\<dots> = 0" using X.prob_space assms by auto
-  finally have X0: "distribution X (X ` space M - {x}) = 0" by auto
-
-  { fix y assume asm: "y \<noteq> x" "y \<in> X ` space M"
-    hence "{y} \<subseteq> X ` space M - {x}" by auto
-    from X.measure_mono[OF this] X0 asm
-    have "distribution X {y} = 0" by auto }
-
-  hence fi: "\<And> y. y \<in> X ` space M \<Longrightarrow> real (distribution X {y}) = (if x = y then 1 else 0)"
-    using assms by auto
-
-  have y: "\<And>y. (if x = y then 1 else 0) * log b (if x = y then 1 else 0) = 0" by simp
-
-  show ?thesis unfolding entropy_eq by (auto simp: y fi)
-qed
-(* --------------- upper bound on entropy for a rv ------------------------- *)
-
-lemma (in finite_prob_space) distribution_1:
-  "distribution X A \<le> 1"
-  unfolding distribution_def measure_space_1[symmetric]
-  by (auto intro!: measure_mono simp: sets_eq_Pow)
-
-lemma (in finite_prob_space) real_distribution_1:
-  "real (distribution X A) \<le> 1"
-  unfolding real_pinfreal_1[symmetric]
-  by (rule real_of_pinfreal_mono[OF _ distribution_1]) simp
+lemma (in finite_information_space) entropy_chain_rule:
+  "\<H>(\<lambda>x. (X x, Y x)) = \<H>(X) + \<H>(Y|X)"
+  unfolding entropy_eq[of X] entropy_eq_cartesian_sum conditional_entropy_eq_cartesian_sum
+  unfolding setsum_commute[of _ "X`space M"] setsum_negf[symmetric] setsum_addf[symmetric]
+  by (rule setsum_cong)
+     (simp_all add: setsum_negf setsum_addf setsum_subtractf setsum_real_distribution
+                    setsum_left_distrib[symmetric] joint_distribution_commute[of X Y])
 
-lemma (in finite_information_space) finite_entropy_le_card:
-  "\<H>(X) \<le> log b (real (card (X ` space M \<inter> {x . distribution X {x} \<noteq> 0})))"
-proof -
-  let "?d x" = "distribution X {x}"
-  let "?p x" = "real (?d x)"
-  have "\<H>(X) = (\<Sum>x\<in>X`space M. ?p x * log b (1 / ?p x))"
-    by (auto intro!: setsum_cong simp: entropy_eq setsum_negf[symmetric])
-  also have "\<dots> \<le> log b (\<Sum>x\<in>X`space M. ?p x * (1 / ?p x))"
-    apply (rule log_setsum')
-    using not_empty b_gt_1 finite_space sum_over_space_real_distribution
-    by auto
-  also have "\<dots> = log b (\<Sum>x\<in>X`space M. if ?d x \<noteq> 0 then 1 else 0)"
-    apply (rule arg_cong[where f="\<lambda>f. log b (\<Sum>x\<in>X`space M. f x)"])
-    using distribution_finite[of X] by (auto simp: expand_fun_eq real_of_pinfreal_eq_0)
-  finally show ?thesis
-    using finite_space by (auto simp: setsum_cases real_eq_of_nat)
-qed
-
-(* --------------- entropy is maximal for a uniform rv --------------------- *)
-
-lemma (in finite_prob_space) uniform_prob:
-  assumes "x \<in> space M"
-  assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
-  shows "prob {x} = 1 / real (card (space M))"
-proof -
-  have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
-    using assms(2)[OF _ `x \<in> space M`] by blast
-  have "1 = prob (space M)"
-    using prob_space by auto
-  also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
-    using real_finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
-      sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
-      finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
-    by (auto simp add:setsum_restrict_set)
-  also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
-    using prob_x by auto
-  also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
-  finally have one: "1 = real (card (space M)) * prob {x}"
-    using real_eq_of_nat by auto
-  hence two: "real (card (space M)) \<noteq> 0" by fastsimp 
-  from one have three: "prob {x} \<noteq> 0" by fastsimp
-  thus ?thesis using one two three divide_cancel_right
-    by (auto simp:field_simps)
-qed
-
-lemma (in finite_information_space) finite_entropy_uniform_max:
-  assumes "\<And>x y. \<lbrakk> x \<in> X ` space M ; y \<in> X ` space M \<rbrakk> \<Longrightarrow> distribution X {x} = distribution X {y}"
-  shows "\<H>(X) = log b (real (card (X ` space M)))"
-proof -
-  note uniform =
-    finite_prob_space_of_images[of X, THEN finite_prob_space.uniform_prob, simplified]
-
-  have card_gt0: "0 < card (X ` space M)" unfolding card_gt_0_iff
-    using finite_space not_empty by auto
-
-  { fix x assume "x \<in> X ` space M"
-    hence "real (distribution X {x}) = 1 / real (card (X ` space M))"
-    proof (rule uniform)
-      fix x y assume "x \<in> X`space M" "y \<in> X`space M"
-      from assms[OF this] show "real (distribution X {x}) = real (distribution X {y})" by simp
-    qed }
-  thus ?thesis
-    using not_empty finite_space b_gt_1 card_gt0
-    by (simp add: entropy_eq real_eq_of_nat[symmetric] log_divide)
-qed
+section {* Partitioning *}
 
 definition "subvimage A f g \<longleftrightarrow> (\<forall>x \<in> A. f -` {f x} \<inter> A \<subseteq> g -` {g x} \<inter> A)"
 
@@ -934,38 +930,6 @@
   "\<H>(f \<circ> X) \<le> \<H>(X)"
   by (subst (2) entropy_partition[of _ "f \<circ> X"]) (auto intro: conditional_entropy_positive)
 
-lemma (in prob_space) distribution_cong:
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
-  shows "distribution X = distribution Y"
-  unfolding distribution_def expand_fun_eq
-  using assms by (auto intro!: arg_cong[where f="\<mu>"])
-
-lemma (in prob_space) joint_distribution_cong:
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
-  assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
-  shows "joint_distribution X Y = joint_distribution X' Y'"
-  unfolding distribution_def expand_fun_eq
-  using assms by (auto intro!: arg_cong[where f="\<mu>"])
-
-lemma image_cong:
-  "\<lbrakk> \<And>x. x \<in> S \<Longrightarrow> X x = X' x \<rbrakk> \<Longrightarrow> X ` S = X' ` S"
-  by (auto intro!: image_eqI)
-
-lemma (in finite_information_space) mutual_information_cong:
-  assumes X: "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
-  assumes Y: "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
-  shows "\<I>(X ; Y) = \<I>(X' ; Y')"
-proof -
-  have "X ` space M = X' ` space M" using X by (rule image_cong)
-  moreover have "Y ` space M = Y' ` space M" using Y by (rule image_cong)
-  ultimately show ?thesis
-  unfolding mutual_information_eq
-    using
-      assms[THEN distribution_cong]
-      joint_distribution_cong[OF assms]
-    by (auto intro!: setsum_cong)
-qed
-
 corollary (in finite_information_space) entropy_of_inj:
   assumes "inj_on f (X`space M)"
   shows "\<H>(f \<circ> X) = \<H>(X)"
--- a/src/HOL/Probability/Lebesgue_Integration.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Lebesgue_Integration.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -209,19 +209,6 @@
     by (auto intro!: **)
 qed
 
-lemma setsum_indicator_disjoint_family:
-  fixes f :: "'d \<Rightarrow> 'e::semiring_1"
-  assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
-  shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
-proof -
-  have "P \<inter> {i. x \<in> A i} = {j}"
-    using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def
-    by auto
-  thus ?thesis
-    unfolding indicator_def
-    by (simp add: if_distrib setsum_cases[OF `finite P`])
-qed
-
 lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence:
   fixes u :: "'a \<Rightarrow> pinfreal"
   assumes u: "u \<in> borel_measurable M"
@@ -426,6 +413,62 @@
   with x show thesis by (auto intro!: that[of f])
 qed
 
+lemma (in sigma_algebra) simple_function_eq_borel_measurable:
+  fixes f :: "'a \<Rightarrow> pinfreal"
+  shows "simple_function f \<longleftrightarrow>
+    finite (f`space M) \<and> f \<in> borel_measurable M"
+  using simple_function_borel_measurable[of f]
+    borel_measurable_simple_function[of f]
+  by (fastsimp simp: simple_function_def)
+
+lemma (in measure_space) simple_function_restricted:
+  fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M"
+  shows "sigma_algebra.simple_function (restricted_space A) f \<longleftrightarrow> simple_function (\<lambda>x. f x * indicator A x)"
+    (is "sigma_algebra.simple_function ?R f \<longleftrightarrow> simple_function ?f")
+proof -
+  interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
+  have "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
+  proof cases
+    assume "A = space M"
+    then have "f`A = ?f`space M" by (fastsimp simp: image_iff)
+    then show ?thesis by simp
+  next
+    assume "A \<noteq> space M"
+    then obtain x where x: "x \<in> space M" "x \<notin> A"
+      using sets_into_space `A \<in> sets M` by auto
+    have *: "?f`space M = f`A \<union> {0}"
+    proof (auto simp add: image_iff)
+      show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0"
+        using x by (auto intro!: bexI[of _ x])
+    next
+      fix x assume "x \<in> A"
+      then show "\<exists>y\<in>space M. f x = f y * indicator A y"
+        using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
+    next
+      fix x
+      assume "indicator A x \<noteq> (0::pinfreal)"
+      then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
+      moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
+      ultimately show "f x = 0" by auto
+    qed
+    then show ?thesis by auto
+  qed
+  then show ?thesis
+    unfolding simple_function_eq_borel_measurable
+      R.simple_function_eq_borel_measurable
+    unfolding borel_measurable_restricted[OF `A \<in> sets M`]
+    by auto
+qed
+
+lemma (in sigma_algebra) simple_function_subalgebra:
+  assumes "sigma_algebra.simple_function (M\<lparr>sets:=N\<rparr>) f"
+  and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets:=N\<rparr>)"
+  shows "simple_function f"
+  using assms
+  unfolding simple_function_def
+  unfolding sigma_algebra.simple_function_def[OF N_subalgebra(2)]
+  by auto
+
 section "Simple integral"
 
 definition (in measure_space)
@@ -668,6 +711,41 @@
   qed
 qed
 
+lemma (in measure_space) simple_integral_restricted:
+  assumes "A \<in> sets M"
+  assumes sf: "simple_function (\<lambda>x. f x * indicator A x)"
+  shows "measure_space.simple_integral (restricted_space A) \<mu> f = simple_integral (\<lambda>x. f x * indicator A x)"
+    (is "_ = simple_integral ?f")
+  unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \<in> sets M`]]
+  unfolding simple_integral_def
+proof (simp, safe intro!: setsum_mono_zero_cong_left)
+  from sf show "finite (?f ` space M)"
+    unfolding simple_function_def by auto
+next
+  fix x assume "x \<in> A"
+  then show "f x \<in> ?f ` space M"
+    using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x])
+next
+  fix x assume "x \<in> space M" "?f x \<notin> f`A"
+  then have "x \<notin> A" by (auto simp: image_iff)
+  then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp
+next
+  fix x assume "x \<in> A"
+  then have "f x \<noteq> 0 \<Longrightarrow>
+    f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M"
+    using `A \<in> sets M` sets_into_space
+    by (auto simp: indicator_def split: split_if_asm)
+  then show "f x * \<mu> (f -` {f x} \<inter> A) =
+    f x * \<mu> (?f -` {f x} \<inter> space M)"
+    unfolding pinfreal_mult_cancel_left by auto
+qed
+
+lemma (in measure_space) simple_integral_subalgebra[simp]:
+  assumes "measure_space (M\<lparr>sets := N\<rparr>) \<mu>"
+  shows "measure_space.simple_integral (M\<lparr>sets := N\<rparr>) \<mu> = simple_integral"
+  unfolding simple_integral_def_raw
+  unfolding measure_space.simple_integral_def_raw[OF assms] by simp
+
 section "Continuous posititve integration"
 
 definition (in measure_space)
@@ -1077,6 +1155,43 @@
   qed
 qed
 
+lemma (in measure_space) positive_integral_translated_density:
+  assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
+  shows "measure_space.positive_integral M (\<lambda>A. positive_integral (\<lambda>x. f x * indicator A x)) g =
+    positive_integral (\<lambda>x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _")
+proof -
+  from measure_space_density[OF assms(1)]
+  interpret T: measure_space M ?T .
+  from borel_measurable_implies_simple_function_sequence[OF assms(2)]
+  obtain G where G: "\<And>i. simple_function (G i)" "G \<up> g" by blast
+  note G_borel = borel_measurable_simple_function[OF this(1)]
+  from T.positive_integral_isoton[OF `G \<up> g` G_borel]
+  have *: "(\<lambda>i. T.positive_integral (G i)) \<up> T.positive_integral g" .
+  { fix i
+    have [simp]: "finite (G i ` space M)"
+      using G(1) unfolding simple_function_def by auto
+    have "T.positive_integral (G i) = T.simple_integral (G i)"
+      using G T.positive_integral_eq_simple_integral by simp
+    also have "\<dots> = positive_integral (\<lambda>x. f x * (\<Sum>y\<in>G i`space M. y * indicator (G i -` {y} \<inter> space M) x))"
+      apply (simp add: T.simple_integral_def)
+      apply (subst positive_integral_cmult[symmetric])
+      using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
+      apply (subst positive_integral_setsum[symmetric])
+      using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
+      by (simp add: setsum_right_distrib field_simps)
+    also have "\<dots> = positive_integral (\<lambda>x. f x * G i x)"
+      by (auto intro!: positive_integral_cong
+               simp: indicator_def if_distrib setsum_cases)
+    finally have "T.positive_integral (G i) = positive_integral (\<lambda>x. f x * G i x)" . }
+  with * have eq_Tg: "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> T.positive_integral g" by simp
+  from G(2) have "(\<lambda>i x. f x * G i x) \<up> (\<lambda>x. f x * g x)"
+    unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right)
+  then have "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> positive_integral (\<lambda>x. f x * g x)"
+    using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times)
+  with eq_Tg show "T.positive_integral g = positive_integral (\<lambda>x. f x * g x)"
+    unfolding isoton_def by simp
+qed
+
 lemma (in measure_space) positive_integral_null_set:
   assumes borel: "u \<in> borel_measurable M" and "N \<in> null_sets"
   shows "positive_integral (\<lambda>x. u x * indicator N x) = 0" (is "?I = 0")
@@ -1222,6 +1337,58 @@
   finally show ?thesis by simp
 qed
 
+lemma (in measure_space) positive_integral_restricted:
+  assumes "A \<in> sets M"
+  shows "measure_space.positive_integral (restricted_space A) \<mu> f = positive_integral (\<lambda>x. f x * indicator A x)"
+    (is "measure_space.positive_integral ?R \<mu> f = positive_integral ?f")
+proof -
+  have msR: "measure_space ?R \<mu>" by (rule restricted_measure_space[OF `A \<in> sets M`])
+  then interpret R: measure_space ?R \<mu> .
+  have saR: "sigma_algebra ?R" by fact
+  have *: "R.positive_integral f = R.positive_integral ?f"
+    by (auto intro!: R.positive_integral_cong)
+  show ?thesis
+    unfolding * R.positive_integral_def positive_integral_def
+    unfolding simple_function_restricted[OF `A \<in> sets M`]
+    apply (simp add: SUPR_def)
+    apply (rule arg_cong[where f=Sup])
+  proof (auto simp: image_iff simple_integral_restricted[OF `A \<in> sets M`])
+    fix g assume "simple_function (\<lambda>x. g x * indicator A x)"
+      "g \<le> f" "\<forall>x\<in>A. \<omega> \<noteq> g x"
+    then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and> (\<forall>y\<in>space M. \<omega> \<noteq> x y) \<and>
+      simple_integral (\<lambda>x. g x * indicator A x) = simple_integral x"
+      apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
+      by (auto simp: indicator_def le_fun_def)
+  next
+    fix g assume g: "simple_function g" "g \<le> (\<lambda>x. f x * indicator A x)"
+      "\<forall>x\<in>space M. \<omega> \<noteq> g x"
+    then have *: "(\<lambda>x. g x * indicator A x) = g"
+      "\<And>x. g x * indicator A x = g x"
+      "\<And>x. g x \<le> f x"
+      by (auto simp: le_fun_def expand_fun_eq indicator_def split: split_if_asm)
+    from g show "\<exists>x. simple_function (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and> (\<forall>xa\<in>A. \<omega> \<noteq> x xa) \<and>
+      simple_integral g = simple_integral (\<lambda>xa. x xa * indicator A xa)"
+      using `A \<in> sets M`[THEN sets_into_space]
+      apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
+      by (fastsimp simp: le_fun_def *)
+  qed
+qed
+
+lemma (in measure_space) positive_integral_subalgebra[simp]:
+  assumes borel: "f \<in> borel_measurable (M\<lparr>sets := N\<rparr>)"
+  and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets := N\<rparr>)"
+  shows "measure_space.positive_integral (M\<lparr>sets := N\<rparr>) \<mu> f = positive_integral f"
+proof -
+  note msN = measure_space_subalgebra[OF N_subalgebra]
+  then interpret N: measure_space "M\<lparr>sets:=N\<rparr>" \<mu> .
+  from N.borel_measurable_implies_simple_function_sequence[OF borel]
+  obtain fs where Nsf: "\<And>i. N.simple_function (fs i)" and "fs \<up> f" by blast
+  then have sf: "\<And>i. simple_function (fs i)"
+    using simple_function_subalgebra[OF _ N_subalgebra] by blast
+  from positive_integral_isoton_simple[OF `fs \<up> f` sf] N.positive_integral_isoton_simple[OF `fs \<up> f` Nsf]
+  show ?thesis unfolding simple_integral_subalgebra[OF msN] isoton_def by simp
+qed
+
 section "Lebesgue Integral"
 
 definition (in measure_space) integrable where
@@ -1629,44 +1796,6 @@
     by (simp add: real_of_pinfreal_eq_0)
 qed
 
-lemma LIMSEQ_max:
-  "u ----> (x::real) \<Longrightarrow> (\<lambda>i. max (u i) 0) ----> max x 0"
-  by (fastsimp intro!: LIMSEQ_I dest!: LIMSEQ_D)
-
-lemma (in sigma_algebra) borel_measurable_LIMSEQ:
-  fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> real"
-  assumes u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x"
-  and u: "\<And>i. u i \<in> borel_measurable M"
-  shows "u' \<in> borel_measurable M"
-proof -
-  let "?pu x i" = "max (u i x) 0"
-  let "?nu x i" = "max (- u i x) 0"
-
-  { fix x assume x: "x \<in> space M"
-    have "(?pu x) ----> max (u' x) 0"
-      "(?nu x) ----> max (- u' x) 0"
-      using u'[OF x] by (auto intro!: LIMSEQ_max LIMSEQ_minus)
-    from LIMSEQ_imp_lim_INF[OF _ this(1)] LIMSEQ_imp_lim_INF[OF _ this(2)]
-    have "(SUP n. INF m. Real (u (n + m) x)) = Real (u' x)"
-      "(SUP n. INF m. Real (- u (n + m) x)) = Real (- u' x)"
-      by (simp_all add: Real_max'[symmetric]) }
-  note eq = this
-
-  have *: "\<And>x. real (Real (u' x)) - real (Real (- u' x)) = u' x"
-    by auto
-
-  have "(SUP n. INF m. (\<lambda>x. Real (u (n + m) x))) \<in> borel_measurable M"
-       "(SUP n. INF m. (\<lambda>x. Real (- u (n + m) x))) \<in> borel_measurable M"
-    using u by (auto intro: borel_measurable_SUP borel_measurable_INF borel_measurable_Real)
-  with eq[THEN measurable_cong, of M "\<lambda>x. x" borel_space]
-  have "(\<lambda>x. Real (u' x)) \<in> borel_measurable M"
-       "(\<lambda>x. Real (- u' x)) \<in> borel_measurable M"
-    unfolding SUPR_fun_expand INFI_fun_expand by auto
-  note this[THEN borel_measurable_real]
-  from borel_measurable_diff[OF this]
-  show ?thesis unfolding * .
-qed
-
 lemma (in measure_space) integral_dominated_convergence:
   assumes u: "\<And>i. integrable (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x"
   and w: "integrable w" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> w x"
@@ -1926,41 +2055,11 @@
     by (simp_all add: integral_cmul_indicator borel_measurable_vimage)
 qed
 
-lemma sigma_algebra_cong:
-  fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
-  assumes *: "sigma_algebra M"
-  and cong: "space M = space M'" "sets M = sets M'"
-  shows "sigma_algebra M'"
-using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .
-
-lemma finite_Pow_additivity_sufficient:
-  assumes "finite (space M)" and "sets M = Pow (space M)"
-  and "positive \<mu>" and "additive M \<mu>"
-  and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
-  shows "finite_measure_space M \<mu>"
-proof -
-  have "sigma_algebra M"
-    using assms by (auto intro!: sigma_algebra_cong[OF sigma_algebra_Pow])
-
-  have "measure_space M \<mu>"
-    by (rule sigma_algebra.finite_additivity_sufficient) (fact+)
-  thus ?thesis
-    unfolding finite_measure_space_def finite_measure_space_axioms_def
-    using assms by simp
-qed
-
-lemma finite_measure_spaceI:
-  assumes "measure_space M \<mu>" and "finite (space M)" and "sets M = Pow (space M)"
-  and "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
-  shows "finite_measure_space M \<mu>"
-  unfolding finite_measure_space_def finite_measure_space_axioms_def
-  using assms by simp
+lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function f"
+  unfolding simple_function_def sets_eq_Pow using finite_space by auto
 
 lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \<in> borel_measurable M"
-  unfolding measurable_def sets_eq_Pow by auto
-
-lemma (in finite_measure_space) simple_function_finite: "simple_function f"
-  unfolding simple_function_def sets_eq_Pow using finite_space by auto
+  by (auto intro: borel_measurable_simple_function)
 
 lemma (in finite_measure_space) positive_integral_finite_eq_setsum:
   "positive_integral f = (\<Sum>x \<in> space M. f x * \<mu> {x})"
@@ -1979,10 +2078,8 @@
     "positive_integral (\<lambda>x. Real (f x)) = (\<Sum>x \<in> space M. Real (f x) * \<mu> {x})"
     "positive_integral (\<lambda>x. Real (- f x)) = (\<Sum>x \<in> space M. Real (- f x) * \<mu> {x})"
     unfolding positive_integral_finite_eq_setsum by auto
-
   show "integrable f" using finite_space finite_measure
     by (simp add: setsum_\<omega> integrable_def sets_eq_Pow)
-
   show ?I using finite_measure
     apply (simp add: integral_def sets_eq_Pow real_of_pinfreal_setsum[symmetric]
       real_of_pinfreal_mult[symmetric] setsum_subtractf[symmetric])
--- a/src/HOL/Probability/Measure.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Measure.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -414,6 +414,19 @@
   finally show ?thesis .
 qed
 
+lemma (in measure_space) measure_finitely_subadditive:
+  assumes "finite I" "A ` I \<subseteq> sets M"
+  shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))"
+using assms proof induct
+  case (insert i I)
+  then have "(\<Union>i\<in>I. A i) \<in> sets M" by (auto intro: finite_UN)
+  then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)"
+    using insert by (simp add: measure_subadditive)
+  also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))"
+    using insert by (auto intro!: add_left_mono)
+  finally show ?case .
+qed simp
+
 lemma (in measure_space) measurable_countably_subadditive:
   assumes "range f \<subseteq> sets M"
   shows "\<mu> (\<Union>i. f i) \<le> (\<Sum>\<^isub>\<infinity> i. \<mu> (f i))"
@@ -432,9 +445,34 @@
   finally show ?thesis .
 qed
 
+lemma (in measure_space) measure_inter_full_set:
+  assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>"
+  assumes T: "\<mu> T = \<mu> (space M)"
+  shows "\<mu> (S \<inter> T) = \<mu> S"
+proof (rule antisym)
+  show " \<mu> (S \<inter> T) \<le> \<mu> S"
+    using assms by (auto intro!: measure_mono)
+
+  show "\<mu> S \<le> \<mu> (S \<inter> T)"
+  proof (rule ccontr)
+    assume contr: "\<not> ?thesis"
+    have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))"
+      unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"])
+    also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)"
+      using assms by (auto intro!: measure_subadditive)
+    also have "\<dots> < \<mu> (T - S) + \<mu> S"
+      by (rule pinfreal_less_add[OF not_\<omega>]) (insert contr, auto)
+    also have "\<dots> = \<mu> (T \<union> S)"
+      using assms by (subst measure_additive) auto
+    also have "\<dots> \<le> \<mu> (space M)"
+      using assms sets_into_space by (auto intro!: measure_mono)
+    finally show False ..
+  qed
+qed
+
 lemma (in measure_space) restricted_measure_space:
   assumes "S \<in> sets M"
-  shows "measure_space (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
+  shows "measure_space (restricted_space S) \<mu>"
     (is "measure_space ?r \<mu>")
   unfolding measure_space_def measure_space_axioms_def
 proof safe
@@ -451,6 +489,46 @@
   qed
 qed
 
+lemma (in measure_space) measure_space_vimage:
+  assumes "f \<in> measurable M M'"
+  and "sigma_algebra M'"
+  shows "measure_space M' (\<lambda>A. \<mu> (f -` A \<inter> space M))" (is "measure_space M' ?T")
+proof -
+  interpret M': sigma_algebra M' by fact
+
+  show ?thesis
+  proof
+    show "?T {} = 0" by simp
+
+    show "countably_additive M' ?T"
+    proof (unfold countably_additive_def, safe)
+      fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets M'" "disjoint_family A"
+      hence *: "\<And>i. f -` (A i) \<inter> space M \<in> sets M"
+        using `f \<in> measurable M M'` by (auto simp: measurable_def)
+      moreover have "(\<Union>i. f -`  A i \<inter> space M) \<in> sets M"
+        using * by blast
+      moreover have **: "disjoint_family (\<lambda>i. f -` A i \<inter> space M)"
+        using `disjoint_family A` by (auto simp: disjoint_family_on_def)
+      ultimately show "(\<Sum>\<^isub>\<infinity> i. ?T (A i)) = ?T (\<Union>i. A i)"
+        using measure_countably_additive[OF _ **] by (auto simp: comp_def vimage_UN)
+    qed
+  qed
+qed
+
+lemma (in measure_space) measure_space_subalgebra:
+  assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
+  shows "measure_space (M\<lparr> sets := N \<rparr>) \<mu>"
+proof -
+  interpret N: sigma_algebra "M\<lparr> sets := N \<rparr>" by fact
+  show ?thesis
+  proof
+    show "countably_additive (M\<lparr>sets := N\<rparr>) \<mu>"
+      using `N \<subseteq> sets M`
+      by (auto simp add: countably_additive_def
+               intro!: measure_countably_additive)
+  qed simp
+qed
+
 section "@{text \<sigma>}-finite Measures"
 
 locale sigma_finite_measure = measure_space +
@@ -458,7 +536,7 @@
 
 lemma (in sigma_finite_measure) restricted_sigma_finite_measure:
   assumes "S \<in> sets M"
-  shows "sigma_finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
+  shows "sigma_finite_measure (restricted_space S) \<mu>"
     (is "sigma_finite_measure ?r _")
   unfolding sigma_finite_measure_def sigma_finite_measure_axioms_def
 proof safe
@@ -486,6 +564,25 @@
   qed
 qed
 
+lemma (in sigma_finite_measure) disjoint_sigma_finite:
+  "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and>
+    (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A"
+proof -
+  obtain A :: "nat \<Rightarrow> 'a set" where
+    range: "range A \<subseteq> sets M" and
+    space: "(\<Union>i. A i) = space M" and
+    measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
+    using sigma_finite by auto
+  note range' = range_disjointed_sets[OF range] range
+  { fix i
+    have "\<mu> (disjointed A i) \<le> \<mu> (A i)"
+      using range' disjointed_subset[of A i] by (auto intro!: measure_mono)
+    then have "\<mu> (disjointed A i) \<noteq> \<omega>"
+      using measure[of i] by auto }
+  with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
+  show ?thesis by (auto intro!: exI[of _ "disjointed A"])
+qed
+
 section "Real measure values"
 
 lemma (in measure_space) real_measure_Union:
@@ -604,7 +701,7 @@
     using finite_measure_of_space by (auto intro!: exI[of _ "\<lambda>x. space M"])
 qed
 
-lemma (in finite_measure) finite_measure:
+lemma (in finite_measure) finite_measure[simp, intro]:
   assumes "A \<in> sets M"
   shows "\<mu> A \<noteq> \<omega>"
 proof -
@@ -619,7 +716,7 @@
 
 lemma (in finite_measure) restricted_finite_measure:
   assumes "S \<in> sets M"
-  shows "finite_measure (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>) \<mu>"
+  shows "finite_measure (restricted_space S) \<mu>"
     (is "finite_measure ?r _")
   unfolding finite_measure_def finite_measure_axioms_def
 proof (safe del: notI)
@@ -707,6 +804,13 @@
   shows "\<mu> (space M - s) = \<mu> (space M) - \<mu> s"
   using measure_compl[OF s, OF finite_measure, OF s] .
 
+lemma (in finite_measure) finite_measure_inter_full_set:
+  assumes "S \<in> sets M" "T \<in> sets M"
+  assumes T: "\<mu> T = \<mu> (space M)"
+  shows "\<mu> (S \<inter> T) = \<mu> S"
+  using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms
+  by auto
+
 section {* Measure preserving *}
 
 definition "measure_preserving A \<mu> B \<nu> =
@@ -817,10 +921,51 @@
   and sets_eq_Pow: "sets M = Pow (space M)"
   and finite_single_measure: "\<And>x. x \<in> space M \<Longrightarrow> \<mu> {x} \<noteq> \<omega>"
 
+lemma (in finite_measure_space) sets_image_space_eq_Pow:
+  "sets (image_space X) = Pow (space (image_space X))"
+proof safe
+  fix x S assume "S \<in> sets (image_space X)" "x \<in> S"
+  then show "x \<in> space (image_space X)"
+    using sets_into_space by (auto intro!: imageI simp: image_space_def)
+next
+  fix S assume "S \<subseteq> space (image_space X)"
+  then obtain S' where "S = X`S'" "S'\<in>sets M"
+    by (auto simp: subset_image_iff sets_eq_Pow image_space_def)
+  then show "S \<in> sets (image_space X)"
+    by (auto simp: image_space_def)
+qed
+
 lemma (in finite_measure_space) sum_over_space: "(\<Sum>x\<in>space M. \<mu> {x}) = \<mu> (space M)"
   using measure_finitely_additive''[of "space M" "\<lambda>i. {i}"]
   by (simp add: sets_eq_Pow disjoint_family_on_def finite_space)
 
+lemma finite_measure_spaceI:
+  assumes "finite (space M)" "sets M = Pow(space M)" and space: "\<mu> (space M) \<noteq> \<omega>"
+    and add: "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B"
+    and "\<mu> {} = 0"
+  shows "finite_measure_space M \<mu>"
+    unfolding finite_measure_space_def finite_measure_space_axioms_def
+proof (safe del: notI)
+  show "measure_space M \<mu>"
+  proof (rule sigma_algebra.finite_additivity_sufficient)
+    show "sigma_algebra M"
+      apply (rule sigma_algebra_cong)
+      apply (rule sigma_algebra_Pow[of "space M"])
+      using assms by simp_all
+    show "finite (space M)" by fact
+    show "positive \<mu>" unfolding positive_def by fact
+    show "additive M \<mu>" unfolding additive_def using assms by simp
+  qed
+  show "finite (space M)" by fact
+  { fix A x assume "A \<in> sets M" "x \<in> A" then show "x \<in> space M"
+      using assms by auto }
+  { fix A assume "A \<subseteq> space M" then show "A \<in> sets M"
+      using assms by auto }
+  { fix x assume *: "x \<in> space M"
+    with add[of "{x}" "space M - {x}"] space
+    show "\<mu> {x} \<noteq> \<omega>" by (auto simp: insert_absorb[OF *] Diff_subset) }
+qed
+
 sublocale finite_measure_space < finite_measure
 proof
   show "\<mu> (space M) \<noteq> \<omega>"
@@ -828,6 +973,22 @@
     using finite_space finite_single_measure by auto
 qed
 
+lemma finite_measure_space_iff:
+  "finite_measure_space M \<mu> \<longleftrightarrow>
+    finite (space M) \<and> sets M = Pow(space M) \<and> \<mu> (space M) \<noteq> \<omega> \<and> \<mu> {} = 0 \<and>
+    (\<forall>A\<subseteq>space M. \<forall>B\<subseteq>space M. A \<inter> B = {} \<longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B)"
+    (is "_ = ?rhs")
+proof (intro iffI)
+  assume "finite_measure_space M \<mu>"
+  then interpret finite_measure_space M \<mu> .
+  show ?rhs
+    using finite_space sets_eq_Pow measure_additive empty_measure finite_measure
+    by auto
+next
+  assume ?rhs then show "finite_measure_space M \<mu>"
+    by (auto intro!: finite_measure_spaceI)
+qed
+
 lemma psuminf_cmult_indicator:
   assumes "disjoint_family A" "x \<in> A i"
   shows "(\<Sum>\<^isub>\<infinity> n. f n * indicator (A n) x) = f i"
--- a/src/HOL/Probability/Positive_Infinite_Real.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Positive_Infinite_Real.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -411,6 +411,10 @@
 lemma pinfreal_less_\<omega>: "x < \<omega> \<longleftrightarrow> x \<noteq> \<omega>"
   by (cases x) auto
 
+lemma pinfreal_0_less_mult_iff[simp]:
+  fixes x y :: pinfreal shows "0 < x * y \<longleftrightarrow> 0 < x \<and> 0 < y"
+  by (cases x, cases y) (auto simp: zero_less_mult_iff)
+
 subsection {* @{text "x - y"} on @{typ pinfreal} *}
 
 instantiation pinfreal :: minus
--- a/src/HOL/Probability/Probability_Space.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Probability_Space.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -1,39 +1,7 @@
 theory Probability_Space
-imports Lebesgue_Integration
+imports Lebesgue_Integration Radon_Nikodym
 begin
 
-lemma (in measure_space) measure_inter_full_set:
-  assumes "S \<in> sets M" "T \<in> sets M" and not_\<omega>: "\<mu> (T - S) \<noteq> \<omega>"
-  assumes T: "\<mu> T = \<mu> (space M)"
-  shows "\<mu> (S \<inter> T) = \<mu> S"
-proof (rule antisym)
-  show " \<mu> (S \<inter> T) \<le> \<mu> S"
-    using assms by (auto intro!: measure_mono)
-
-  show "\<mu> S \<le> \<mu> (S \<inter> T)"
-  proof (rule ccontr)
-    assume contr: "\<not> ?thesis"
-    have "\<mu> (space M) = \<mu> ((T - S) \<union> (S \<inter> T))"
-      unfolding T[symmetric] by (auto intro!: arg_cong[where f="\<mu>"])
-    also have "\<dots> \<le> \<mu> (T - S) + \<mu> (S \<inter> T)"
-      using assms by (auto intro!: measure_subadditive)
-    also have "\<dots> < \<mu> (T - S) + \<mu> S"
-      by (rule pinfreal_less_add[OF not_\<omega>]) (insert contr, auto)
-    also have "\<dots> = \<mu> (T \<union> S)"
-      using assms by (subst measure_additive) auto
-    also have "\<dots> \<le> \<mu> (space M)"
-      using assms sets_into_space by (auto intro!: measure_mono)
-    finally show False ..
-  qed
-qed
-
-lemma (in finite_measure) finite_measure_inter_full_set:
-  assumes "S \<in> sets M" "T \<in> sets M"
-  assumes T: "\<mu> T = \<mu> (space M)"
-  shows "\<mu> (S \<inter> T) = \<mu> S"
-  using measure_inter_full_set[OF assms(1,2) finite_measure assms(3)] assms
-  by auto
-
 locale prob_space = measure_space +
   assumes measure_space_1: "\<mu> (space M) = 1"
 
@@ -63,6 +31,19 @@
 abbreviation
   "joint_distribution X Y \<equiv> distribution (\<lambda>x. (X x, Y x))"
 
+lemma (in prob_space) distribution_cong:
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = Y x"
+  shows "distribution X = distribution Y"
+  unfolding distribution_def expand_fun_eq
+  using assms by (auto intro!: arg_cong[where f="\<mu>"])
+
+lemma (in prob_space) joint_distribution_cong:
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> X x = X' x"
+  assumes "\<And>x. x \<in> space M \<Longrightarrow> Y x = Y' x"
+  shows "joint_distribution X Y = joint_distribution X' Y'"
+  unfolding distribution_def expand_fun_eq
+  using assms by (auto intro!: arg_cong[where f="\<mu>"])
+
 lemma prob_space: "prob (space M) = 1"
   unfolding measure_space_1 by simp
 
@@ -75,10 +56,6 @@
   finally show ?thesis .
 qed
 
-lemma measure_finite[simp, intro]:
-  assumes "A \<in> events" shows "\<mu> A \<noteq> \<omega>"
-  using measure_le_1[OF assms] by auto
-
 lemma prob_compl:
   assumes "A \<in> events"
   shows "prob (space M - A) = 1 - prob A"
@@ -197,35 +174,17 @@
 qed
 
 lemma distribution_prob_space:
-  fixes S :: "('c, 'd) algebra_scheme"
-  assumes "sigma_algebra S" "random_variable S X"
+  assumes S: "sigma_algebra S" "random_variable S X"
   shows "prob_space S (distribution X)"
 proof -
-  interpret S: sigma_algebra S by fact
+  interpret S: measure_space S "distribution X"
+    using measure_space_vimage[OF S(2,1)] unfolding distribution_def .
   show ?thesis
   proof
-    show "distribution X {} = 0" unfolding distribution_def by simp
     have "X -` space S \<inter> space M = space M"
       using `random_variable S X` by (auto simp: measurable_def)
-    then show "distribution X (space S) = 1" using measure_space_1 by (simp add: distribution_def)
-
-    show "countably_additive S (distribution X)"
-    proof (unfold countably_additive_def, safe)
-      fix A :: "nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets S" "disjoint_family A"
-      hence *: "\<And>i. X -` A i \<inter> space M \<in> sets M"
-        using `random_variable S X` by (auto simp: measurable_def)
-      moreover hence "\<And>i. \<mu> (X -` A i \<inter> space M) \<noteq> \<omega>"
-        using finite_measure by auto
-      moreover have "(\<Union>i. X -`  A i \<inter> space M) \<in> sets M"
-        using * by blast
-      moreover hence "\<mu> (\<Union>i. X -` A i \<inter> space M) \<noteq> \<omega>"
-        using finite_measure by auto
-      moreover have **: "disjoint_family (\<lambda>i. X -` A i \<inter> space M)"
-        using `disjoint_family A` by (auto simp: disjoint_family_on_def)
-      ultimately show "(\<Sum>\<^isub>\<infinity> i. distribution X (A i)) = distribution X (\<Union>i. A i)"
-        using measure_countably_additive[OF _ **]
-        by (auto simp: distribution_def Real_real comp_def vimage_UN)
-    qed
+    then show "distribution X (space S) = 1"
+      using measure_space_1 by (simp add: distribution_def)
   qed
 qed
 
@@ -379,39 +338,246 @@
     joint_distribution_restriction_snd[of X Y "{(x, y)}"]
   by auto
 
+lemma (in finite_prob_space) distribution_mono:
+  assumes "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
+  shows "distribution X x \<le> distribution Y y"
+  unfolding distribution_def
+  using assms by (auto simp: sets_eq_Pow intro!: measure_mono)
+
+lemma (in finite_prob_space) distribution_mono_gt_0:
+  assumes gt_0: "0 < distribution X x"
+  assumes *: "\<And>t. \<lbrakk> t \<in> space M ; X t \<in> x \<rbrakk> \<Longrightarrow> Y t \<in> y"
+  shows "0 < distribution Y y"
+  by (rule less_le_trans[OF gt_0 distribution_mono]) (rule *)
+
+lemma (in finite_prob_space) sum_over_space_distrib:
+  "(\<Sum>x\<in>X`space M. distribution X {x}) = 1"
+  unfolding distribution_def measure_space_1[symmetric] using finite_space
+  by (subst measure_finitely_additive'')
+     (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=\<mu>])
+
+lemma (in finite_prob_space) sum_over_space_real_distribution:
+  "(\<Sum>x\<in>X`space M. real (distribution X {x})) = 1"
+  unfolding distribution_def prob_space[symmetric] using finite_space
+  by (subst real_finite_measure_finite_Union[symmetric])
+     (auto simp add: disjoint_family_on_def sets_eq_Pow intro!: arg_cong[where f=prob])
+
+lemma (in finite_prob_space) finite_sum_over_space_eq_1:
+  "(\<Sum>x\<in>space M. real (\<mu> {x})) = 1"
+  using sum_over_space_eq_1 finite_measure by (simp add: real_of_pinfreal_setsum sets_eq_Pow)
+
+lemma (in finite_prob_space) distribution_finite:
+  "distribution X A \<noteq> \<omega>"
+  using finite_measure[of "X -` A \<inter> space M"]
+  unfolding distribution_def sets_eq_Pow by auto
+
+lemma (in finite_prob_space) real_distribution_gt_0[simp]:
+  "0 < real (distribution Y y) \<longleftrightarrow>  0 < distribution Y y"
+  using assms by (auto intro!: real_pinfreal_pos distribution_finite)
+
+lemma (in finite_prob_space) real_distribution_mult_pos_pos:
+  assumes "0 < distribution Y y"
+  and "0 < distribution X x"
+  shows "0 < real (distribution Y y * distribution X x)"
+  unfolding real_of_pinfreal_mult[symmetric]
+  using assms by (auto intro!: mult_pos_pos)
+
+lemma (in finite_prob_space) real_distribution_divide_pos_pos:
+  assumes "0 < distribution Y y"
+  and "0 < distribution X x"
+  shows "0 < real (distribution Y y / distribution X x)"
+  unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric]
+  using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
+
+lemma (in finite_prob_space) real_distribution_mult_inverse_pos_pos:
+  assumes "0 < distribution Y y"
+  and "0 < distribution X x"
+  shows "0 < real (distribution Y y * inverse (distribution X x))"
+  unfolding divide_pinfreal_def real_of_pinfreal_mult[symmetric]
+  using assms distribution_finite[of X x] by (cases "distribution X x") (auto intro!: mult_pos_pos)
+
+lemma (in prob_space) distribution_remove_const:
+  shows "joint_distribution X (\<lambda>x. ()) {(x, ())} = distribution X {x}"
+  and "joint_distribution (\<lambda>x. ()) X {((), x)} = distribution X {x}"
+  and "joint_distribution X (\<lambda>x. (Y x, ())) {(x, y, ())} = joint_distribution X Y {(x, y)}"
+  and "joint_distribution X (\<lambda>x. ((), Y x)) {(x, (), y)} = joint_distribution X Y {(x, y)}"
+  and "distribution (\<lambda>x. ()) {()} = 1"
+  unfolding measure_space_1[symmetric]
+  by (auto intro!: arg_cong[where f="\<mu>"] simp: distribution_def)
+
+lemma (in finite_prob_space) setsum_distribution_gen:
+  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
+  and "inj_on f (X`space M)"
+  shows "(\<Sum>x \<in> X`space M. distribution Y {f x}) = distribution Z {c}"
+  unfolding distribution_def assms
+  using finite_space assms
+  by (subst measure_finitely_additive'')
+     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
+      intro!: arg_cong[where f=prob])
+
+lemma (in finite_prob_space) setsum_distribution:
+  "(\<Sum>x \<in> X`space M. joint_distribution X Y {(x, y)}) = distribution Y {y}"
+  "(\<Sum>y \<in> Y`space M. joint_distribution X Y {(x, y)}) = distribution X {x}"
+  "(\<Sum>x \<in> X`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution Y Z {(y, z)}"
+  "(\<Sum>y \<in> Y`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Z {(x, z)}"
+  "(\<Sum>z \<in> Z`space M. joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)}) = joint_distribution X Y {(x, y)}"
+  by (auto intro!: inj_onI setsum_distribution_gen)
+
+lemma (in finite_prob_space) setsum_real_distribution_gen:
+  assumes "Z -` {c} \<inter> space M = (\<Union>x \<in> X`space M. Y -` {f x}) \<inter> space M"
+  and "inj_on f (X`space M)"
+  shows "(\<Sum>x \<in> X`space M. real (distribution Y {f x})) = real (distribution Z {c})"
+  unfolding distribution_def assms
+  using finite_space assms
+  by (subst real_finite_measure_finite_Union[symmetric])
+     (auto simp add: disjoint_family_on_def sets_eq_Pow inj_on_def
+        intro!: arg_cong[where f=prob])
+
+lemma (in finite_prob_space) setsum_real_distribution:
+  "(\<Sum>x \<in> X`space M. real (joint_distribution X Y {(x, y)})) = real (distribution Y {y})"
+  "(\<Sum>y \<in> Y`space M. real (joint_distribution X Y {(x, y)})) = real (distribution X {x})"
+  "(\<Sum>x \<in> X`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution Y Z {(y, z)})"
+  "(\<Sum>y \<in> Y`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Z {(x, z)})"
+  "(\<Sum>z \<in> Z`space M. real (joint_distribution X (\<lambda>x. (Y x, Z x)) {(x, y, z)})) = real (joint_distribution X Y {(x, y)})"
+  by (auto intro!: inj_onI setsum_real_distribution_gen)
+
+lemma (in finite_prob_space) real_distribution_order:
+  shows "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution X {x})"
+  and "r \<le> real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r \<le> real (distribution Y {y})"
+  and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution X {x})"
+  and "r < real (joint_distribution X Y {(x, y)}) \<Longrightarrow> r < real (distribution Y {y})"
+  and "distribution X {x} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
+  and "distribution Y {y} = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
+  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
+  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
+  using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"]
+  by auto
+
+lemma (in prob_space) joint_distribution_remove[simp]:
+    "joint_distribution X X {(x, x)} = distribution X {x}"
+  unfolding distribution_def by (auto intro!: arg_cong[where f="\<mu>"])
+
+lemma (in finite_prob_space) distribution_1:
+  "distribution X A \<le> 1"
+  unfolding distribution_def measure_space_1[symmetric]
+  by (auto intro!: measure_mono simp: sets_eq_Pow)
+
+lemma (in finite_prob_space) real_distribution_1:
+  "real (distribution X A) \<le> 1"
+  unfolding real_pinfreal_1[symmetric]
+  by (rule real_of_pinfreal_mono[OF _ distribution_1]) simp
+
+lemma (in finite_prob_space) uniform_prob:
+  assumes "x \<in> space M"
+  assumes "\<And> x y. \<lbrakk>x \<in> space M ; y \<in> space M\<rbrakk> \<Longrightarrow> prob {x} = prob {y}"
+  shows "prob {x} = 1 / real (card (space M))"
+proof -
+  have prob_x: "\<And> y. y \<in> space M \<Longrightarrow> prob {y} = prob {x}"
+    using assms(2)[OF _ `x \<in> space M`] by blast
+  have "1 = prob (space M)"
+    using prob_space by auto
+  also have "\<dots> = (\<Sum> x \<in> space M. prob {x})"
+    using real_finite_measure_finite_Union[of "space M" "\<lambda> x. {x}", simplified]
+      sets_eq_Pow inj_singleton[unfolded inj_on_def, rule_format]
+      finite_space unfolding disjoint_family_on_def  prob_space[symmetric]
+    by (auto simp add:setsum_restrict_set)
+  also have "\<dots> = (\<Sum> y \<in> space M. prob {x})"
+    using prob_x by auto
+  also have "\<dots> = real_of_nat (card (space M)) * prob {x}" by simp
+  finally have one: "1 = real (card (space M)) * prob {x}"
+    using real_eq_of_nat by auto
+  hence two: "real (card (space M)) \<noteq> 0" by fastsimp 
+  from one have three: "prob {x} \<noteq> 0" by fastsimp
+  thus ?thesis using one two three divide_cancel_right
+    by (auto simp:field_simps)
+qed
+
+lemma (in prob_space) prob_space_subalgebra:
+  assumes "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
+  shows "prob_space (M\<lparr> sets := N \<rparr>) \<mu>"
+proof -
+  interpret N: measure_space "M\<lparr> sets := N \<rparr>" \<mu>
+    using measure_space_subalgebra[OF assms] .
+  show ?thesis
+    proof qed (simp add: measure_space_1)
+qed
+
+lemma (in prob_space) prob_space_of_restricted_space:
+  assumes "\<mu> A \<noteq> 0" "\<mu> A \<noteq> \<omega>" "A \<in> sets M"
+  shows "prob_space (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
+  unfolding prob_space_def prob_space_axioms_def
+proof
+  show "\<mu> (space (restricted_space A)) / \<mu> A = 1"
+    using `\<mu> A \<noteq> 0` `\<mu> A \<noteq> \<omega>` by (auto simp: pinfreal_noteq_omega_Ex)
+  have *: "\<And>S. \<mu> S / \<mu> A = inverse (\<mu> A) * \<mu> S" by (simp add: mult_commute)
+  interpret A: measure_space "restricted_space A" \<mu>
+    using `A \<in> sets M` by (rule restricted_measure_space)
+  show "measure_space (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
+  proof
+    show "\<mu> {} / \<mu> A = 0" by auto
+    show "countably_additive (restricted_space A) (\<lambda>S. \<mu> S / \<mu> A)"
+        unfolding countably_additive_def psuminf_cmult_right *
+        using A.measure_countably_additive by auto
+  qed
+qed
+
+lemma finite_prob_spaceI:
+  assumes "finite (space M)" "sets M = Pow(space M)" "\<mu> (space M) = 1" "\<mu> {} = 0"
+    and "\<And>A B. A\<subseteq>space M \<Longrightarrow> B\<subseteq>space M \<Longrightarrow> A \<inter> B = {} \<Longrightarrow> \<mu> (A \<union> B) = \<mu> A + \<mu> B"
+  shows "finite_prob_space M \<mu>"
+  unfolding finite_prob_space_eq
+proof
+  show "finite_measure_space M \<mu>" using assms
+     by (auto intro!: finite_measure_spaceI)
+  show "\<mu> (space M) = 1" by fact
+qed
+
+lemma (in finite_prob_space) finite_measure_space:
+  fixes X :: "'a \<Rightarrow> 'x"
+  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
+    (is "finite_measure_space ?S _")
+proof (rule finite_measure_spaceI, simp_all)
+  show "finite (X ` space M)" using finite_space by simp
+next
+  fix A B :: "'x set" assume "A \<inter> B = {}"
+  then show "distribution X (A \<union> B) = distribution X A + distribution X B"
+    unfolding distribution_def
+    by (subst measure_additive)
+       (auto intro!: arg_cong[where f=\<mu>] simp: sets_eq_Pow)
+qed
+
+lemma (in finite_prob_space) finite_prob_space_of_images:
+  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
+  by (simp add: finite_prob_space_eq finite_measure_space)
+
+lemma (in prob_space) joint_distribution_commute:
+  "joint_distribution X Y x = joint_distribution Y X ((\<lambda>(x,y). (y,x))`x)"
+  unfolding distribution_def by (auto intro!: arg_cong[where f=\<mu>])
+
+lemma (in finite_prob_space) real_distribution_order':
+  shows "real (distribution X {x}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
+  and "real (distribution Y {y}) = 0 \<Longrightarrow> real (joint_distribution X Y {(x, y)}) = 0"
+  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_fst, of X Y "{(x, y)}"]
+  using real_of_pinfreal_mono[OF distribution_finite joint_distribution_restriction_snd, of X Y "{(x, y)}"]
+  using real_pinfreal_nonneg[of "joint_distribution X Y {(x, y)}"]
+  by auto
+
 lemma (in finite_prob_space) finite_product_measure_space:
+  fixes X :: "'a \<Rightarrow> 'x" and Y :: "'a \<Rightarrow> 'y"
   assumes "finite s1" "finite s2"
   shows "finite_measure_space \<lparr> space = s1 \<times> s2, sets = Pow (s1 \<times> s2)\<rparr> (joint_distribution X Y)"
     (is "finite_measure_space ?M ?D")
-proof (rule finite_Pow_additivity_sufficient)
-  show "positive ?D"
-    unfolding positive_def using assms sets_eq_Pow
-    by (simp add: distribution_def)
-
-  show "additive ?M ?D" unfolding additive_def
-  proof safe
-    fix x y
-    have A: "((\<lambda>x. (X x, Y x)) -` x) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto
-    have B: "((\<lambda>x. (X x, Y x)) -` y) \<inter> space M \<in> sets M" using assms sets_eq_Pow by auto
-    assume "x \<inter> y = {}"
-    hence "(\<lambda>x. (X x, Y x)) -` x \<inter> space M \<inter> ((\<lambda>x. (X x, Y x)) -` y \<inter> space M) = {}"
-      by auto
-    from additive[unfolded additive_def, rule_format, OF A B] this
-      finite_measure[OF A] finite_measure[OF B]
-    show "?D (x \<union> y) = ?D x + ?D y"
-      apply (simp add: distribution_def)
-      apply (subst Int_Un_distrib2)
-      by (auto simp: real_of_pinfreal_add)
-  qed
-
-  show "finite (space ?M)"
+proof (rule finite_measure_spaceI, simp_all)
+  show "finite (s1 \<times> s2)"
     using assms by auto
-
-  show "sets ?M = Pow (space ?M)"
-    by simp
-
-  { fix x assume "x \<in> space ?M" thus "?D {x} \<noteq> \<omega>"
-    unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) }
+  show "joint_distribution X Y (s1\<times>s2) \<noteq> \<omega>"
+    using distribution_finite .
+next
+  fix A B :: "('x*'y) set" assume "A \<inter> B = {}"
+  then show "joint_distribution X Y (A \<union> B) = joint_distribution X Y A + joint_distribution X Y B"
+    unfolding distribution_def
+    by (subst measure_additive)
+       (auto intro!: arg_cong[where f=\<mu>] simp: sets_eq_Pow)
 qed
 
 lemma (in finite_prob_space) finite_product_measure_space_of_images:
@@ -420,47 +586,133 @@
                               (joint_distribution X Y)"
   using finite_space by (auto intro!: finite_product_measure_space)
 
-lemma (in finite_prob_space) finite_measure_space:
-  shows "finite_measure_space \<lparr>space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
-    (is "finite_measure_space ?S _")
-proof (rule finite_Pow_additivity_sufficient, simp_all)
-  show "finite (X ` space M)" using finite_space by simp
+section "Conditional Expectation and Probability"
 
-  show "positive (distribution X)"
-    unfolding distribution_def positive_def using sets_eq_Pow by auto
+lemma (in prob_space) conditional_expectation_exists:
+  fixes X :: "'a \<Rightarrow> pinfreal"
+  assumes borel: "X \<in> borel_measurable M"
+  and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
+  shows "\<exists>Y\<in>borel_measurable (M\<lparr> sets := N \<rparr>). \<forall>C\<in>N.
+      positive_integral (\<lambda>x. Y x * indicator C x) = positive_integral (\<lambda>x. X x * indicator C x)"
+proof -
+  interpret P: prob_space "M\<lparr> sets := N \<rparr>" \<mu>
+    using prob_space_subalgebra[OF N_subalgebra] .
+
+  let "?f A" = "\<lambda>x. X x * indicator A x"
+  let "?Q A" = "positive_integral (?f A)"
+
+  from measure_space_density[OF borel]
+  have Q: "measure_space (M\<lparr> sets := N \<rparr>) ?Q"
+    by (rule measure_space.measure_space_subalgebra[OF _ N_subalgebra])
+  then interpret Q: measure_space "M\<lparr> sets := N \<rparr>" ?Q .
 
-  show "additive ?S (distribution X)" unfolding additive_def distribution_def
-  proof (simp, safe)
-    fix x y
-    have x: "(X -` x) \<inter> space M \<in> sets M"
-      and y: "(X -` y) \<inter> space M \<in> sets M" using sets_eq_Pow by auto
-    assume "x \<inter> y = {}"
-    hence "X -` x \<inter> space M \<inter> (X -` y \<inter> space M) = {}" by auto
-    from additive[unfolded additive_def, rule_format, OF x y] this
-      finite_measure[OF x] finite_measure[OF y]
-    have "\<mu> (((X -` x) \<union> (X -` y)) \<inter> space M) =
-      \<mu> ((X -` x) \<inter> space M) + \<mu> ((X -` y) \<inter> space M)"
-      by (subst Int_Un_distrib2) auto
-    thus "\<mu> ((X -` x \<union> X -` y) \<inter> space M) = \<mu> (X -` x \<inter> space M) + \<mu> (X -` y \<inter> space M)"
-      by auto
+  have "P.absolutely_continuous ?Q"
+    unfolding P.absolutely_continuous_def
+  proof (safe, simp)
+    fix A assume "A \<in> N" "\<mu> A = 0"
+    moreover then have f_borel: "?f A \<in> borel_measurable M"
+      using borel N_subalgebra by (auto intro: borel_measurable_indicator)
+    moreover have "{x\<in>space M. ?f A x \<noteq> 0} = (?f A -` {0<..} \<inter> space M) \<inter> A"
+      by (auto simp: indicator_def)
+    moreover have "\<mu> \<dots> \<le> \<mu> A"
+      using `A \<in> N` N_subalgebra f_borel
+      by (auto intro!: measure_mono Int[of _ A] measurable_sets)
+    ultimately show "?Q A = 0"
+      by (simp add: positive_integral_0_iff)
   qed
-
-  { fix x assume "x \<in> X ` space M" thus "distribution X {x} \<noteq> \<omega>"
-    unfolding distribution_def by (auto intro!: finite_measure simp: sets_eq_Pow) }
+  from P.Radon_Nikodym[OF Q this]
+  obtain Y where Y: "Y \<in> borel_measurable (M\<lparr>sets := N\<rparr>)"
+    "\<And>A. A \<in> sets (M\<lparr>sets:=N\<rparr>) \<Longrightarrow> ?Q A = P.positive_integral (\<lambda>x. Y x * indicator A x)"
+    by blast
+  with N_subalgebra show ?thesis
+    by (auto intro!: bexI[OF _ Y(1)])
 qed
 
-lemma (in finite_prob_space) finite_prob_space_of_images:
-  "finite_prob_space \<lparr> space = X ` space M, sets = Pow (X ` space M)\<rparr> (distribution X)"
-  by (simp add: finite_prob_space_eq finite_measure_space)
+definition (in prob_space)
+  "conditional_expectation N X = (SOME Y. Y\<in>borel_measurable (M\<lparr>sets:=N\<rparr>)
+    \<and> (\<forall>C\<in>N. positive_integral (\<lambda>x. Y x * indicator C x) = positive_integral (\<lambda>x. X x * indicator C x)))"
+
+abbreviation (in prob_space)
+  "conditional_prob N A \<equiv> conditional_expectation N (indicator A)"
+
+lemma (in prob_space)
+  fixes X :: "'a \<Rightarrow> pinfreal"
+  assumes borel: "X \<in> borel_measurable M"
+  and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr> sets := N \<rparr>)"
+  shows borel_measurable_conditional_expectation:
+    "conditional_expectation N X \<in> borel_measurable (M\<lparr> sets := N \<rparr>)"
+  and conditional_expectation: "\<And>C. C \<in> N \<Longrightarrow>
+      positive_integral (\<lambda>x. conditional_expectation N X x * indicator C x) =
+      positive_integral (\<lambda>x. X x * indicator C x)"
+   (is "\<And>C. C \<in> N \<Longrightarrow> ?eq C")
+proof -
+  note CE = conditional_expectation_exists[OF assms, unfolded Bex_def]
+  then show "conditional_expectation N X \<in> borel_measurable (M\<lparr> sets := N \<rparr>)"
+    unfolding conditional_expectation_def by (rule someI2_ex) blast
+
+  from CE show "\<And>C. C\<in>N \<Longrightarrow> ?eq C"
+    unfolding conditional_expectation_def by (rule someI2_ex) blast
+qed
+
+lemma (in sigma_algebra) factorize_measurable_function:
+  fixes Z :: "'a \<Rightarrow> pinfreal" and Y :: "'a \<Rightarrow> 'c"
+  assumes "sigma_algebra M'" and "Y \<in> measurable M M'" "Z \<in> borel_measurable M"
+  shows "Z \<in> borel_measurable (sigma_algebra.vimage_algebra M' (space M) Y)
+    \<longleftrightarrow> (\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x))"
+proof safe
+  interpret M': sigma_algebra M' by fact
+  have Y: "Y \<in> space M \<rightarrow> space M'" using assms unfolding measurable_def by auto
+  from M'.sigma_algebra_vimage[OF this]
+  interpret va: sigma_algebra "M'.vimage_algebra (space M) Y" .
 
-lemma (in finite_prob_space) finite_product_prob_space_of_images:
-  "finite_prob_space \<lparr> space = X ` space M \<times> Y ` space M, sets = Pow (X ` space M \<times> Y ` space M)\<rparr>
-                     (joint_distribution X Y)"
-  (is "finite_prob_space ?S _")
-proof (simp add: finite_prob_space_eq finite_product_measure_space_of_images)
-  have "X -` X ` space M \<inter> Y -` Y ` space M \<inter> space M = space M" by auto
-  thus "joint_distribution X Y (X ` space M \<times> Y ` space M) = 1"
-    by (simp add: distribution_def prob_space vimage_Times comp_def measure_space_1)
+  { fix g :: "'c \<Rightarrow> pinfreal" assume "g \<in> borel_measurable M'"
+    with M'.measurable_vimage_algebra[OF Y]
+    have "g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+      by (rule measurable_comp)
+    moreover assume "\<forall>x\<in>space M. Z x = g (Y x)"
+    then have "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y) \<longleftrightarrow>
+       g \<circ> Y \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+       by (auto intro!: measurable_cong)
+    ultimately show "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+      by simp }
+
+  assume "Z \<in> borel_measurable (M'.vimage_algebra (space M) Y)"
+  from va.borel_measurable_implies_simple_function_sequence[OF this]
+  obtain f where f: "\<And>i. va.simple_function (f i)" and "f \<up> Z" by blast
+
+  have "\<forall>i. \<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
+  proof
+    fix i
+    from f[of i] have "finite (f i`space M)" and B_ex:
+      "\<forall>z\<in>(f i)`space M. \<exists>B. B \<in> sets M' \<and> (f i) -` {z} \<inter> space M = Y -` B \<inter> space M"
+      unfolding va.simple_function_def by auto
+    from B_ex[THEN bchoice] guess B .. note B = this
+
+    let ?g = "\<lambda>x. \<Sum>z\<in>f i`space M. z * indicator (B z) x"
+
+    show "\<exists>g. M'.simple_function g \<and> (\<forall>x\<in>space M. f i x = g (Y x))"
+    proof (intro exI[of _ ?g] conjI ballI)
+      show "M'.simple_function ?g" using B by auto
+
+      fix x assume "x \<in> space M"
+      then have "\<And>z. z \<in> f i`space M \<Longrightarrow> indicator (B z) (Y x) = (indicator (f i -` {z} \<inter> space M) x::pinfreal)"
+        unfolding indicator_def using B by auto
+      then show "f i x = ?g (Y x)" using `x \<in> space M` f[of i]
+        by (subst va.simple_function_indicator_representation) auto
+    qed
+  qed
+  from choice[OF this] guess g .. note g = this
+
+  show "\<exists>g\<in>borel_measurable M'. \<forall>x\<in>space M. Z x = g (Y x)"
+  proof (intro ballI bexI)
+    show "(SUP i. g i) \<in> borel_measurable M'"
+      using g by (auto intro: M'.borel_measurable_simple_function)
+    fix x assume "x \<in> space M"
+    have "Z x = (SUP i. f i) x" using `f \<up> Z` unfolding isoton_def by simp
+    also have "\<dots> = (SUP i. g i) (Y x)" unfolding SUPR_fun_expand
+      using g `x \<in> space M` by simp
+    finally show "Z x = (SUP i. g i) (Y x)" .
+  qed
 qed
 
 end
--- a/src/HOL/Probability/Product_Measure.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Product_Measure.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -2,14 +2,412 @@
 imports Lebesgue_Integration
 begin
 
+definition "dynkin M \<longleftrightarrow>
+  space M \<in> sets M \<and>
+  (\<forall> A \<in> sets M. A \<subseteq> space M) \<and>
+  (\<forall> a \<in> sets M. \<forall> b \<in> sets M. a \<subseteq> b \<longrightarrow> b - a \<in> sets M) \<and>
+  (\<forall>A. disjoint_family A \<and> range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)"
+
+lemma dynkinI:
+  assumes "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M"
+  assumes "space M \<in> sets M" and "\<forall> b \<in> sets M. \<forall> a \<in> sets M. a \<subseteq> b \<longrightarrow> b - a \<in> sets M"
+  assumes "\<And> a. (\<And> i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {})
+          \<Longrightarrow> (\<And> i :: nat. a i \<in> sets M) \<Longrightarrow> UNION UNIV a \<in> sets M"
+  shows "dynkin M"
+using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff)
+
+lemma dynkin_subset:
+  assumes "dynkin M"
+  shows "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M"
+using assms unfolding dynkin_def by auto
+
+lemma dynkin_space:
+  assumes "dynkin M"
+  shows "space M \<in> sets M"
+using assms unfolding dynkin_def by auto
+
+lemma dynkin_diff:
+  assumes "dynkin M"
+  shows "\<And> a b. \<lbrakk> a \<in> sets M ; b \<in> sets M ; a \<subseteq> b \<rbrakk> \<Longrightarrow> b - a \<in> sets M"
+using assms unfolding dynkin_def by auto
+
+lemma dynkin_empty:
+  assumes "dynkin M"
+  shows "{} \<in> sets M"
+using dynkin_diff[OF assms dynkin_space[OF assms] dynkin_space[OF assms]] by auto
+
+lemma dynkin_UN:
+  assumes "dynkin M"
+  assumes "\<And> i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}"
+  assumes "\<And> i :: nat. a i \<in> sets M"
+  shows "UNION UNIV a \<in> sets M"
+using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff)
+
+definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)"
+
+lemma dynkin_trivial:
+  shows "dynkin \<lparr> space = A, sets = Pow A \<rparr>"
+by (rule dynkinI) auto
+
+lemma dynkin_lemma:
+  fixes D :: "'a algebra"
+  assumes stab: "Int_stable E"
+  and spac: "space E = space D"
+  and subsED: "sets E \<subseteq> sets D"
+  and subsDE: "sets D \<subseteq> sigma_sets (space E) (sets E)"
+  and dyn: "dynkin D"
+  shows "sigma (space E) (sets E) = D"
+proof -
+  def sets_\<delta>E == "\<Inter> {sets d | d :: 'a algebra. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}"
+  def \<delta>E == "\<lparr> space = space E, sets = sets_\<delta>E \<rparr>"
+  have "\<lparr> space = space E, sets = Pow (space E) \<rparr> \<in> {d | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}"
+    using dynkin_trivial spac subsED dynkin_subset[OF dyn] by fastsimp
+  hence not_empty: "{sets (d :: 'a algebra) | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d} \<noteq> {}"
+    using exI[of "\<lambda> x. space x = space E \<and> dynkin x \<and> sets E \<subseteq> sets x" "\<lparr> space = space E, sets = Pow (space E) \<rparr>", simplified]
+    by auto
+  have \<delta>E_D: "sets_\<delta>E \<subseteq> sets D"
+    unfolding sets_\<delta>E_def using assms by auto
+  have \<delta>ynkin: "dynkin \<delta>E"
+  proof (rule dynkinI, safe)
+    fix A x assume asm: "A \<in> sets \<delta>E" "x \<in> A"
+    { fix d :: "('a, 'b) algebra_scheme" assume "A \<in> sets d" "dynkin d \<and> space d = space E"
+      hence "A \<subseteq> space d" using dynkin_subset by auto }
+    show "x \<in> space \<delta>E" using asm unfolding \<delta>E_def sets_\<delta>E_def using not_empty
+      by simp (metis dynkin_subset in_mono mem_def)
+  next
+    show "space \<delta>E \<in> sets \<delta>E"
+      unfolding \<delta>E_def sets_\<delta>E_def
+      using dynkin_space by fastsimp
+  next
+    fix a b assume "a \<in> sets \<delta>E" "b \<in> sets \<delta>E" "a \<subseteq> b"
+    thus "b - a \<in> sets \<delta>E"
+      unfolding \<delta>E_def sets_\<delta>E_def by (auto intro:dynkin_diff)
+  next
+    fix a assume asm: "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" "\<And>i. a i \<in> sets \<delta>E"
+    thus "UNION UNIV a \<in> sets \<delta>E"
+      unfolding \<delta>E_def sets_\<delta>E_def apply (auto intro!:dynkin_UN[OF _ asm(1)])
+      by blast
+  qed
+
+  def Dy == "\<lambda> d. {A | A. A \<in> sets_\<delta>E \<and> A \<inter> d \<in> sets_\<delta>E}"
+  { fix d assume dasm: "d \<in> sets_\<delta>E"
+    have "dynkin \<lparr> space = space E, sets = Dy d \<rparr>"
+    proof (rule dynkinI, safe, simp_all)
+      fix A x assume "A \<in> Dy d" "x \<in> A"
+      thus "x \<in> space E" unfolding Dy_def sets_\<delta>E_def using not_empty
+        by simp (metis dynkin_subset in_mono mem_def)
+    next
+      show "space E \<in> Dy d"
+        unfolding Dy_def \<delta>E_def sets_\<delta>E_def
+      proof auto
+        fix d assume asm: "dynkin d" "space d = space E" "sets E \<subseteq> sets d"
+        hence "space d \<in> sets d" using dynkin_space[OF asm(1)] by auto
+        thus "space E \<in> sets d" using asm by auto
+      next
+        fix da :: "'a algebra" assume asm: "dynkin da" "space da = space E" "sets E \<subseteq> sets da"
+        have d: "d = space E \<inter> d"
+          using dasm dynkin_subset[OF asm(1)] asm(2) dynkin_subset[OF \<delta>ynkin]
+          unfolding \<delta>E_def by auto
+        hence "space E \<inter> d \<in> sets \<delta>E" unfolding \<delta>E_def
+          using dasm by auto
+        have "sets \<delta>E \<subseteq> sets da" unfolding \<delta>E_def sets_\<delta>E_def using asm
+          by auto
+        thus "space E \<inter> d \<in> sets da" using dasm asm d dynkin_subset[OF \<delta>ynkin]
+          unfolding \<delta>E_def by auto
+      qed
+    next
+      fix a b assume absm: "a \<in> Dy d" "b \<in> Dy d" "a \<subseteq> b"
+      hence "a \<in> sets \<delta>E" "b \<in> sets \<delta>E"
+        unfolding Dy_def \<delta>E_def by auto
+      hence *: "b - a \<in> sets \<delta>E"
+        using dynkin_diff[OF \<delta>ynkin] `a \<subseteq> b` by auto
+      have "a \<inter> d \<in> sets \<delta>E" "b \<inter> d \<in> sets \<delta>E"
+        using absm unfolding Dy_def \<delta>E_def by auto
+      hence "(b \<inter> d) - (a \<inter> d) \<in> sets \<delta>E"
+        using dynkin_diff[OF \<delta>ynkin] `a \<subseteq> b` by auto
+      hence **: "(b - a) \<inter> d \<in> sets \<delta>E" by (auto simp add:Diff_Int_distrib2)
+      thus "b - a \<in> Dy d"
+        using * ** unfolding Dy_def \<delta>E_def by auto
+    next
+      fix a assume aasm: "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" "\<And>i. a i \<in> Dy d"
+      hence "\<And> i. a i \<in> sets \<delta>E"
+        unfolding Dy_def \<delta>E_def by auto
+      from dynkin_UN[OF \<delta>ynkin aasm(1) this]
+      have *: "UNION UNIV a \<in> sets \<delta>E" by auto
+      from aasm
+      have aE: "\<forall> i. a i \<inter> d \<in> sets \<delta>E"
+        unfolding Dy_def \<delta>E_def by auto
+      from aasm
+      have "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> (a i \<inter> d) \<inter> (a j \<inter> d) = {}" by auto
+      from dynkin_UN[OF \<delta>ynkin this]
+      have "UNION UNIV (\<lambda> i. a i \<inter> d) \<in> sets \<delta>E"
+        using aE by auto
+      hence **: "UNION UNIV a \<inter> d \<in> sets \<delta>E" by auto
+      from * ** show "UNION UNIV a \<in> Dy d" unfolding Dy_def \<delta>E_def by auto
+    qed } note Dy_nkin = this
+  have E_\<delta>E: "sets E \<subseteq> sets \<delta>E"
+    unfolding \<delta>E_def sets_\<delta>E_def by auto
+  { fix d assume dasm: "d \<in> sets \<delta>E"
+    { fix e assume easm: "e \<in> sets E"
+      hence deasm: "e \<in> sets \<delta>E"
+        unfolding \<delta>E_def sets_\<delta>E_def by auto
+      have subset: "Dy e \<subseteq> sets \<delta>E"
+        unfolding Dy_def \<delta>E_def by auto
+      { fix e' assume e'asm: "e' \<in> sets E"
+        have "e' \<inter> e \<in> sets E"
+          using easm e'asm stab unfolding Int_stable_def by auto
+        hence "e' \<inter> e \<in> sets \<delta>E"
+          unfolding \<delta>E_def sets_\<delta>E_def by auto
+        hence "e' \<in> Dy e" using e'asm unfolding Dy_def \<delta>E_def sets_\<delta>E_def by auto }
+      hence E_Dy: "sets E \<subseteq> Dy e" by auto
+      have "\<lparr> space = space E, sets = Dy e \<rparr> \<in> {d | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}"
+        using Dy_nkin[OF deasm[unfolded \<delta>E_def, simplified]] E_\<delta>E E_Dy by auto
+      hence "sets_\<delta>E \<subseteq> Dy e"
+        unfolding sets_\<delta>E_def by auto (metis E_Dy simps(1) simps(2) spac)
+      hence "sets \<delta>E = Dy e" using subset unfolding \<delta>E_def by auto
+      hence "d \<inter> e \<in> sets \<delta>E"
+        using dasm easm deasm unfolding Dy_def \<delta>E_def by auto
+      hence "e \<in> Dy d" using deasm
+        unfolding Dy_def \<delta>E_def
+        by (auto simp add:Int_commute) }
+    hence "sets E \<subseteq> Dy d" by auto
+    hence "sets \<delta>E \<subseteq> Dy d" using Dy_nkin[OF dasm[unfolded \<delta>E_def, simplified]]
+      unfolding \<delta>E_def sets_\<delta>E_def
+      by auto (metis `sets E <= Dy d` simps(1) simps(2) spac)
+    hence *: "sets \<delta>E = Dy d"
+      unfolding Dy_def \<delta>E_def by auto
+    fix a assume aasm: "a \<in> sets \<delta>E"
+    hence "a \<inter> d \<in> sets \<delta>E"
+      using * dasm unfolding Dy_def \<delta>E_def by auto } note \<delta>E_stab = this
+  { fix A :: "nat \<Rightarrow> 'a set" assume Asm: "range A \<subseteq> sets \<delta>E" "\<And>A. A \<in> sets \<delta>E \<Longrightarrow> A \<subseteq> space \<delta>E"
+      "\<And>a. a \<in> sets \<delta>E \<Longrightarrow> space \<delta>E - a \<in> sets \<delta>E"
+    "{} \<in> sets \<delta>E" "space \<delta>E \<in> sets \<delta>E"
+    let "?A i" = "A i \<inter> (\<Inter> j \<in> {..< i}. space \<delta>E - A j)"
+    { fix i :: nat
+      have *: "(\<Inter> j \<in> {..< i}. space \<delta>E - A j) \<inter> space \<delta>E \<in> sets \<delta>E"
+        apply (induct i)
+        using lessThan_Suc Asm \<delta>E_stab apply fastsimp
+        apply (subst lessThan_Suc)
+        apply (subst INT_insert)
+        apply (subst Int_assoc)
+        apply (subst \<delta>E_stab)
+        using lessThan_Suc Asm \<delta>E_stab Asm
+        apply (fastsimp simp add:Int_assoc dynkin_diff[OF \<delta>ynkin])
+        prefer 2 apply simp
+        apply (rule dynkin_diff[OF \<delta>ynkin, of _ "space \<delta>E", OF _ dynkin_space[OF \<delta>ynkin]])
+        using Asm by auto
+      have **: "\<And> i. A i \<subseteq> space \<delta>E" using Asm by auto
+      have "(\<Inter> j \<in> {..< i}. space \<delta>E - A j) \<subseteq> space \<delta>E \<or> (\<Inter> j \<in> {..< i}. A j) = UNIV \<and> i = 0"
+        apply (cases i)
+        using Asm ** dynkin_subset[OF \<delta>ynkin, of "A (i - 1)"]
+        by auto
+      hence Aisets: "?A i \<in> sets \<delta>E"
+        apply (cases i)
+        using Asm * apply fastsimp
+        apply (rule \<delta>E_stab)
+        using Asm * **
+        by (auto simp add:Int_absorb2)
+      have "?A i = disjointed A i" unfolding disjointed_def
+      atLeast0LessThan using Asm by auto
+      hence "?A i = disjointed A i" "?A i \<in> sets \<delta>E"
+        using Aisets by auto
+    } note Ai = this
+    from dynkin_UN[OF \<delta>ynkin _ this(2)] this disjoint_family_disjointed[of A]
+    have "(\<Union> i. ?A i) \<in> sets \<delta>E"
+      by (auto simp add:disjoint_family_on_def disjointed_def)
+    hence "(\<Union> i. A i) \<in> sets \<delta>E"
+      using Ai(1) UN_disjointed_eq[of A] by auto } note \<delta>E_UN = this
+  { fix a b assume asm: "a \<in> sets \<delta>E" "b \<in> sets \<delta>E"
+    let ?ab = "\<lambda> i. if (i::nat) = 0 then a else if i = 1 then b else {}"
+    have *: "(\<Union> i. ?ab i) \<in> sets \<delta>E"
+      apply (rule \<delta>E_UN)
+      using asm \<delta>E_UN dynkin_empty[OF \<delta>ynkin] 
+      dynkin_subset[OF \<delta>ynkin] 
+      dynkin_space[OF \<delta>ynkin]
+      dynkin_diff[OF \<delta>ynkin] by auto
+    have "(\<Union> i. ?ab i) = a \<union> b" apply auto
+      apply (case_tac "i = 0")
+      apply auto
+      apply (case_tac "i = 1")
+      by auto
+    hence "a \<union> b \<in> sets \<delta>E" using * by auto} note \<delta>E_Un = this
+  have "sigma_algebra \<delta>E"
+    apply unfold_locales
+    using dynkin_subset[OF \<delta>ynkin]
+    using dynkin_diff[OF \<delta>ynkin, of _ "space \<delta>E", OF _ dynkin_space[OF \<delta>ynkin]]
+    using dynkin_diff[OF \<delta>ynkin, of "space \<delta>E" "space \<delta>E", OF dynkin_space[OF \<delta>ynkin] dynkin_space[OF \<delta>ynkin]]
+    using dynkin_space[OF \<delta>ynkin]
+    using \<delta>E_UN \<delta>E_Un
+    by auto
+  from sigma_algebra.sigma_subset[OF this E_\<delta>E] \<delta>E_D subsDE spac
+  show ?thesis by (auto simp add:\<delta>E_def sigma_def)
+qed
+
+lemma measure_eq:
+  assumes fin: "\<mu> (space M) = \<nu> (space M)" "\<nu> (space M) < \<omega>"
+  assumes E: "M = sigma (space E) (sets E)" "Int_stable E"
+  assumes eq: "\<And> e. e \<in> sets E \<Longrightarrow> \<mu> e = \<nu> e"
+  assumes ms: "measure_space M \<mu>" "measure_space M \<nu>"
+  assumes A: "A \<in> sets M"
+  shows "\<mu> A = \<nu> A"
+proof -
+  interpret M: measure_space M \<mu>
+    using ms by simp
+  interpret M': measure_space M \<nu>
+    using ms by simp
+
+  let ?D_sets = "{A. A \<in> sets M \<and> \<mu> A = \<nu> A}"
+  have \<delta>: "dynkin \<lparr> space = space M , sets = ?D_sets \<rparr>"
+  proof (rule dynkinI, safe, simp_all)
+    fix A x assume "A \<in> sets M \<and> \<mu> A = \<nu> A" "x \<in> A"
+    thus "x \<in> space M" using assms M.sets_into_space by auto
+  next
+    show "\<mu> (space M) = \<nu> (space M)"
+      using fin by auto
+  next
+    fix a b
+    assume asm: "a \<in> sets M \<and> \<mu> a = \<nu> a"
+      "b \<in> sets M \<and> \<mu> b = \<nu> b" "a \<subseteq> b"
+    hence "a \<subseteq> space M"
+      using M.sets_into_space by auto
+    from M.measure_mono[OF this]
+    have "\<mu> a \<le> \<mu> (space M)"
+      using asm by auto
+    hence afin: "\<mu> a < \<omega>"
+      using fin by auto
+    have *: "b = b - a \<union> a" using asm by auto
+    have **: "(b - a) \<inter> a = {}" using asm by auto
+    have iv: "\<mu> (b - a) + \<mu> a = \<mu> b"
+      using M.measure_additive[of "b - a" a]
+        conjunct1[OF asm(1)] conjunct1[OF asm(2)] * **
+      by auto
+    have v: "\<nu> (b - a) + \<nu> a = \<nu> b"
+      using M'.measure_additive[of "b - a" a]
+        conjunct1[OF asm(1)] conjunct1[OF asm(2)] * **
+      by auto
+    from iv v have "\<mu> (b - a) = \<nu> (b - a)" using asm afin
+      pinfreal_add_cancel_right[of "\<mu> (b - a)" "\<nu> a" "\<nu> (b - a)"]
+      by auto
+    thus "b - a \<in> sets M \<and> \<mu> (b - a) = \<nu> (b - a)"
+      using asm by auto
+  next
+    fix a assume "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}"
+      "\<And>i. a i \<in> sets M \<and> \<mu> (a i) = \<nu> (a i)"
+    thus "(\<Union>x. a x) \<in> sets M \<and> \<mu> (\<Union>x. a x) = \<nu> (\<Union>x. a x)"
+      using M.measure_countably_additive
+        M'.measure_countably_additive
+        M.countable_UN
+      apply (auto simp add:disjoint_family_on_def image_def)
+      apply (subst M.measure_countably_additive[symmetric])
+      apply (auto simp add:disjoint_family_on_def)
+      apply (subst M'.measure_countably_additive[symmetric])
+      by (auto simp add:disjoint_family_on_def)
+  qed
+  have *: "sets E \<subseteq> ?D_sets"
+    using eq E sigma_sets.Basic[of _ "sets E"]
+    by (auto simp add:sigma_def)
+  have **: "?D_sets \<subseteq> sets M" by auto
+  have "M = \<lparr> space = space M , sets = ?D_sets \<rparr>"
+    unfolding E(1)
+    apply (rule dynkin_lemma[OF E(2)])
+    using eq E space_sigma \<delta> sigma_sets.Basic
+    by (auto simp add:sigma_def)
+  from subst[OF this, of "\<lambda> M. A \<in> sets M", OF A]
+  show ?thesis by auto
+qed
+(*
+lemma
+  assumes sfin: "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And> i :: nat. \<nu> (A i) < \<omega>"
+  assumes A: "\<And>  i. \<mu> (A i) = \<nu> (A i)" "\<And> i. A i \<subseteq> A (Suc i)"
+  assumes E: "M = sigma (space E) (sets E)" "Int_stable E"
+  assumes eq: "\<And> e. e \<in> sets E \<Longrightarrow> \<mu> e = \<nu> e"
+  assumes ms: "measure_space (M :: 'a algebra) \<mu>" "measure_space M \<nu>"
+  assumes B: "B \<in> sets M"
+  shows "\<mu> B = \<nu> B"
+proof -
+  interpret M: measure_space M \<mu> by (rule ms)
+  interpret M': measure_space M \<nu> by (rule ms)
+  have *: "M = \<lparr> space = space M, sets = sets M \<rparr>" by auto
+  { fix i :: nat
+    have **: "M\<lparr> space := A i, sets := op \<inter> (A i) ` sets M \<rparr> =
+      \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr>"
+      by auto
+    have mu_i: "measure_space \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr> \<mu>"
+      using M.restricted_measure_space[of "A i", simplified **]
+        sfin by auto
+    have nu_i: "measure_space \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr> \<nu>"
+      using M'.restricted_measure_space[of "A i", simplified **]
+        sfin by auto
+    let ?M = "\<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr>"
+    have "\<mu> (A i \<inter> B) = \<nu> (A i \<inter> B)"
+      apply (rule measure_eq[of \<mu> ?M \<nu> "\<lparr> space = space E \<inter> A i, sets = op \<inter> (A i) ` sets E\<rparr>" "A i \<inter> B", simplified])
+      using assms nu_i mu_i
+      apply (auto simp add:image_def) (* TODO *) sorry
+    show ?thesis sorry
+qed
+*)
 definition prod_sets where
   "prod_sets A B = {z. \<exists>x \<in> A. \<exists>y \<in> B. z = x \<times> y}"
 
 definition
-  "prod_measure M \<mu> N \<nu> = (\<lambda>A. measure_space.positive_integral M \<mu> (\<lambda>s0. \<nu> ((\<lambda>s1. (s0, s1)) -` A)))"
+  "prod_measure_space M1 M2 = sigma (space M1 \<times> space M2) (prod_sets (sets M1) (sets M2))"
+
+lemma
+  fixes M1 :: "'a algebra" and M2 :: "'b algebra"
+  assumes "algebra M1" "algebra M2"
+  shows measureable_fst[intro!, simp]:
+    "fst \<in> measurable (prod_measure_space M1 M2) M1" (is ?fst)
+  and measureable_snd[intro!, simp]:
+    "snd \<in> measurable (prod_measure_space M1 M2) M2" (is ?snd)
+proof -
+  interpret M1: algebra M1 by fact
+  interpret M2: algebra M2 by fact
+
+  { fix X assume "X \<in> sets M1"
+    then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. fst -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
+      apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"])
+      using M1.sets_into_space by force+ }
+  moreover
+  { fix X assume "X \<in> sets M2"
+    then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. snd -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2"
+      apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X])
+      using M2.sets_into_space by force+ }
+  ultimately show ?fst ?snd
+    by (force intro!: sigma_sets.Basic
+              simp: measurable_def prod_measure_space_def prod_sets_def sets_sigma)+
+qed
+
+lemma (in sigma_algebra) measureable_prod:
+  fixes M1 :: "'a algebra" and M2 :: "'b algebra"
+  assumes "algebra M1" "algebra M2"
+  shows "f \<in> measurable M (prod_measure_space M1 M2) \<longleftrightarrow>
+    (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"
+using assms proof (safe intro!: measurable_comp[where b="prod_measure_space M1 M2"])
+  interpret M1: algebra M1 by fact
+  interpret M2: algebra M2 by fact
+  assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2"
+
+  show "f \<in> measurable M (prod_measure_space M1 M2)" unfolding prod_measure_space_def
+  proof (rule measurable_sigma)
+    show "prod_sets (sets M1) (sets M2) \<subseteq> Pow (space M1 \<times> space M2)"
+      unfolding prod_sets_def using M1.sets_into_space M2.sets_into_space by auto
+    show "f \<in> space M \<rightarrow> space M1 \<times> space M2"
+      using f s by (auto simp: mem_Times_iff measurable_def comp_def)
+    fix A assume "A \<in> prod_sets (sets M1) (sets M2)"
+    then obtain B C where "B \<in> sets M1" "C \<in> sets M2" "A = B \<times> C"
+      unfolding prod_sets_def by auto
+    moreover have "(fst \<circ> f) -` B \<inter> space M \<in> sets M"
+      using f `B \<in> sets M1` unfolding measurable_def by auto
+    moreover have "(snd \<circ> f) -` C \<inter> space M \<in> sets M"
+      using s `C \<in> sets M2` unfolding measurable_def by auto
+    moreover have "f -` A \<inter> space M = ((fst \<circ> f) -` B \<inter> space M) \<inter> ((snd \<circ> f) -` C \<inter> space M)"
+      unfolding `A = B \<times> C` by (auto simp: vimage_Times)
+    ultimately show "f -` A \<inter> space M \<in> sets M" by auto
+  qed
+qed
 
 definition
-  "prod_measure_space M1 M2 = sigma (space M1 \<times> space M2) (prod_sets (sets M1) (sets M2))"
+  "prod_measure M \<mu> N \<nu> = (\<lambda>A. measure_space.positive_integral M \<mu> (\<lambda>s0. \<nu> ((\<lambda>s1. (s0, s1)) -` A)))"
 
 lemma prod_setsI: "x \<in> A \<Longrightarrow> y \<in> B \<Longrightarrow> (x \<times> y) \<in> prod_sets A B"
   by (auto simp add: prod_sets_def)
@@ -114,36 +512,25 @@
 qed
 
 lemma (in finite_measure_space) finite_measure_space_finite_prod_measure:
-  assumes "finite_measure_space N \<nu>"
+  fixes N :: "('c, 'd) algebra_scheme"
+  assumes N: "finite_measure_space N \<nu>"
   shows "finite_measure_space (prod_measure_space M N) (prod_measure M \<mu> N \<nu>)"
   unfolding finite_prod_measure_space[OF assms]
-proof (rule finite_measure_spaceI)
+proof (rule finite_measure_spaceI, simp_all)
   interpret N: finite_measure_space N \<nu> by fact
-
-  let ?P = "\<lparr>space = space M \<times> space N, sets = Pow (space M \<times> space N)\<rparr>"
-  show "measure_space ?P (prod_measure M \<mu> N \<nu>)"
-  proof (rule sigma_algebra.finite_additivity_sufficient)
-    show "sigma_algebra ?P" by (rule sigma_algebra_Pow)
-    show "finite (space ?P)" using finite_space N.finite_space by auto
-    from finite_prod_measure_times[OF assms, of "{}" "{}"]
-    show "positive (prod_measure M \<mu> N \<nu>)"
-      unfolding positive_def by simp
+  show "finite (space M \<times> space N)" using finite_space N.finite_space by auto
+  show "prod_measure M \<mu> N \<nu> (space M \<times> space N) \<noteq> \<omega>"
+    using finite_prod_measure_times[OF N top N.top] by simp
+  show "prod_measure M \<mu> N \<nu> {} = 0"
+    using finite_prod_measure_times[OF N empty_sets N.empty_sets] by simp
 
-    show "additive ?P (prod_measure M \<mu> N \<nu>)"
-      unfolding additive_def
-      apply (auto simp add: sets_eq_Pow prod_measure_def positive_integral_add[symmetric]
-                  intro!: positive_integral_cong)
-      apply (subst N.measure_additive[symmetric])
-      by (auto simp: N.sets_eq_Pow sets_eq_Pow)
-  qed
-  show "finite (space ?P)" using finite_space N.finite_space by auto
-  show "sets ?P = Pow (space ?P)" by simp
-
-  fix x assume "x \<in> space ?P"
-  with finite_prod_measure_times[OF assms, of "{fst x}" "{snd x}"]
-    finite_measure[of "{fst x}"] N.finite_measure[of "{snd x}"]
-  show "prod_measure M \<mu> N \<nu> {x} \<noteq> \<omega>"
-    by (auto simp add: sets_eq_Pow N.sets_eq_Pow elim!: SigmaE)
+  fix A B :: "('a * 'c) set" assume "A \<inter> B = {}" "A \<subseteq> space M \<times> space N" "B \<subseteq> space M \<times> space N"
+  then show "prod_measure M \<mu> N \<nu> (A \<union> B) = prod_measure M \<mu> N \<nu> A + prod_measure M \<mu> N \<nu> B"
+    apply (auto simp add: sets_eq_Pow prod_measure_def positive_integral_add[symmetric]
+                intro!: positive_integral_cong)
+    apply (subst N.measure_additive)
+    apply (auto intro!: arg_cong[where f=\<mu>] simp: N.sets_eq_Pow sets_eq_Pow)
+    done
 qed
 
 lemma (in finite_measure_space) finite_measure_space_finite_prod_measure_alterantive:
@@ -153,4 +540,18 @@
   unfolding finite_prod_measure_space[OF N, symmetric]
   using finite_measure_space_finite_prod_measure[OF N] .
 
-end
\ No newline at end of file
+lemma prod_measure_times_finite:
+  assumes fms: "finite_measure_space M \<mu>" "finite_measure_space N \<nu>" and a: "a \<in> space M \<times> space N"
+  shows "prod_measure M \<mu> N \<nu> {a} = \<mu> {fst a} * \<nu> {snd a}"
+proof (cases a)
+  case (Pair b c)
+  hence a_eq: "{a} = {b} \<times> {c}" by simp
+
+  interpret M: finite_measure_space M \<mu> by fact
+  interpret N: finite_measure_space N \<nu> by fact
+
+  from finite_measure_space.finite_prod_measure_times[OF fms, of "{b}" "{c}"] M.sets_eq_Pow N.sets_eq_Pow a Pair
+  show ?thesis unfolding a_eq by simp
+qed
+
+end
--- a/src/HOL/Probability/Radon_Nikodym.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Radon_Nikodym.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -2,201 +2,6 @@
 imports Lebesgue_Integration
 begin
 
-lemma (in measure_space) measure_finitely_subadditive:
-  assumes "finite I" "A ` I \<subseteq> sets M"
-  shows "\<mu> (\<Union>i\<in>I. A i) \<le> (\<Sum>i\<in>I. \<mu> (A i))"
-using assms proof induct
-  case (insert i I)
-  then have "(\<Union>i\<in>I. A i) \<in> sets M" by (auto intro: finite_UN)
-  then have "\<mu> (\<Union>i\<in>insert i I. A i) \<le> \<mu> (A i) + \<mu> (\<Union>i\<in>I. A i)"
-    using insert by (simp add: measure_subadditive)
-  also have "\<dots> \<le> (\<Sum>i\<in>insert i I. \<mu> (A i))"
-    using insert by (auto intro!: add_left_mono)
-  finally show ?case .
-qed simp
-
-lemma (in sigma_algebra) borel_measurable_restricted:
-  fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M"
-  shows "f \<in> borel_measurable (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<longleftrightarrow>
-    (\<lambda>x. f x * indicator A x) \<in> borel_measurable M"
-    (is "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable M")
-proof -
-  interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
-  have *: "f \<in> borel_measurable ?R \<longleftrightarrow> ?f \<in> borel_measurable ?R"
-    by (auto intro!: measurable_cong)
-  show ?thesis unfolding *
-    unfolding in_borel_measurable_borel_space
-  proof (simp, safe)
-    fix S :: "pinfreal set" assume "S \<in> sets borel_space"
-      "\<forall>S\<in>sets borel_space. ?f -` S \<inter> A \<in> op \<inter> A ` sets M"
-    then have "?f -` S \<inter> A \<in> op \<inter> A ` sets M" by auto
-    then have f: "?f -` S \<inter> A \<in> sets M"
-      using `A \<in> sets M` sets_into_space by fastsimp
-    show "?f -` S \<inter> space M \<in> sets M"
-    proof cases
-      assume "0 \<in> S"
-      then have "?f -` S \<inter> space M = ?f -` S \<inter> A \<union> (space M - A)"
-        using `A \<in> sets M` sets_into_space by auto
-      then show ?thesis using f `A \<in> sets M` by (auto intro!: Un Diff)
-    next
-      assume "0 \<notin> S"
-      then have "?f -` S \<inter> space M = ?f -` S \<inter> A"
-        using `A \<in> sets M` sets_into_space
-        by (auto simp: indicator_def split: split_if_asm)
-      then show ?thesis using f by auto
-    qed
-  next
-    fix S :: "pinfreal set" assume "S \<in> sets borel_space"
-      "\<forall>S\<in>sets borel_space. ?f -` S \<inter> space M \<in> sets M"
-    then have f: "?f -` S \<inter> space M \<in> sets M" by auto
-    then show "?f -` S \<inter> A \<in> op \<inter> A ` sets M"
-      using `A \<in> sets M` sets_into_space
-      apply (simp add: image_iff)
-      apply (rule bexI[OF _ f])
-      by auto
-  qed
-qed
-
-lemma (in sigma_algebra) simple_function_eq_borel_measurable:
-  fixes f :: "'a \<Rightarrow> pinfreal"
-  shows "simple_function f \<longleftrightarrow>
-    finite (f`space M) \<and> f \<in> borel_measurable M"
-  using simple_function_borel_measurable[of f]
-    borel_measurable_simple_function[of f]
-  by (fastsimp simp: simple_function_def)
-
-lemma (in measure_space) simple_function_restricted:
-  fixes f :: "'a \<Rightarrow> pinfreal" assumes "A \<in> sets M"
-  shows "sigma_algebra.simple_function (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) f \<longleftrightarrow> simple_function (\<lambda>x. f x * indicator A x)"
-    (is "sigma_algebra.simple_function ?R f \<longleftrightarrow> simple_function ?f")
-proof -
-  interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`])
-  have "finite (f`A) \<longleftrightarrow> finite (?f`space M)"
-  proof cases
-    assume "A = space M"
-    then have "f`A = ?f`space M" by (fastsimp simp: image_iff)
-    then show ?thesis by simp
-  next
-    assume "A \<noteq> space M"
-    then obtain x where x: "x \<in> space M" "x \<notin> A"
-      using sets_into_space `A \<in> sets M` by auto
-    have *: "?f`space M = f`A \<union> {0}"
-    proof (auto simp add: image_iff)
-      show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0"
-        using x by (auto intro!: bexI[of _ x])
-    next
-      fix x assume "x \<in> A"
-      then show "\<exists>y\<in>space M. f x = f y * indicator A y"
-        using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x])
-    next
-      fix x
-      assume "indicator A x \<noteq> (0::pinfreal)"
-      then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm)
-      moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y"
-      ultimately show "f x = 0" by auto
-    qed
-    then show ?thesis by auto
-  qed
-  then show ?thesis
-    unfolding simple_function_eq_borel_measurable
-      R.simple_function_eq_borel_measurable
-    unfolding borel_measurable_restricted[OF `A \<in> sets M`]
-    by auto
-qed
-
-lemma (in measure_space) simple_integral_restricted:
-  assumes "A \<in> sets M"
-  assumes sf: "simple_function (\<lambda>x. f x * indicator A x)"
-  shows "measure_space.simple_integral (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<mu> f = simple_integral (\<lambda>x. f x * indicator A x)"
-    (is "_ = simple_integral ?f")
-  unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \<in> sets M`]]
-  unfolding simple_integral_def
-proof (simp, safe intro!: setsum_mono_zero_cong_left)
-  from sf show "finite (?f ` space M)"
-    unfolding simple_function_def by auto
-next
-  fix x assume "x \<in> A"
-  then show "f x \<in> ?f ` space M"
-    using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x])
-next
-  fix x assume "x \<in> space M" "?f x \<notin> f`A"
-  then have "x \<notin> A" by (auto simp: image_iff)
-  then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp
-next
-  fix x assume "x \<in> A"
-  then have "f x \<noteq> 0 \<Longrightarrow>
-    f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M"
-    using `A \<in> sets M` sets_into_space
-    by (auto simp: indicator_def split: split_if_asm)
-  then show "f x * \<mu> (f -` {f x} \<inter> A) =
-    f x * \<mu> (?f -` {f x} \<inter> space M)"
-    unfolding pinfreal_mult_cancel_left by auto
-qed
-
-lemma (in measure_space) positive_integral_restricted:
-  assumes "A \<in> sets M"
-  shows "measure_space.positive_integral (M\<lparr> space := A, sets := op \<inter> A ` sets M \<rparr>) \<mu> f = positive_integral (\<lambda>x. f x * indicator A x)"
-    (is "measure_space.positive_integral ?R \<mu> f = positive_integral ?f")
-proof -
-  have msR: "measure_space ?R \<mu>" by (rule restricted_measure_space[OF `A \<in> sets M`])
-  then interpret R: measure_space ?R \<mu> .
-  have saR: "sigma_algebra ?R" by fact
-  have *: "R.positive_integral f = R.positive_integral ?f"
-    by (auto intro!: R.positive_integral_cong)
-  show ?thesis
-    unfolding * R.positive_integral_def positive_integral_def
-    unfolding simple_function_restricted[OF `A \<in> sets M`]
-    apply (simp add: SUPR_def)
-    apply (rule arg_cong[where f=Sup])
-  proof (auto simp: image_iff simple_integral_restricted[OF `A \<in> sets M`])
-    fix g assume "simple_function (\<lambda>x. g x * indicator A x)"
-      "g \<le> f" "\<forall>x\<in>A. \<omega> \<noteq> g x"
-    then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and> (\<forall>y\<in>space M. \<omega> \<noteq> x y) \<and>
-      simple_integral (\<lambda>x. g x * indicator A x) = simple_integral x"
-      apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
-      by (auto simp: indicator_def le_fun_def)
-  next
-    fix g assume g: "simple_function g" "g \<le> (\<lambda>x. f x * indicator A x)"
-      "\<forall>x\<in>space M. \<omega> \<noteq> g x"
-    then have *: "(\<lambda>x. g x * indicator A x) = g"
-      "\<And>x. g x * indicator A x = g x"
-      "\<And>x. g x \<le> f x"
-      by (auto simp: le_fun_def expand_fun_eq indicator_def split: split_if_asm)
-    from g show "\<exists>x. simple_function (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and> (\<forall>xa\<in>A. \<omega> \<noteq> x xa) \<and>
-      simple_integral g = simple_integral (\<lambda>xa. x xa * indicator A xa)"
-      using `A \<in> sets M`[THEN sets_into_space]
-      apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"])
-      by (fastsimp simp: le_fun_def *)
-  qed
-qed
-
-lemma (in sigma_algebra) borel_measurable_psuminf:
-  assumes "\<And>i. f i \<in> borel_measurable M"
-  shows "(\<lambda>x. (\<Sum>\<^isub>\<infinity> i. f i x)) \<in> borel_measurable M"
-  using assms unfolding psuminf_def
-  by (auto intro!: borel_measurable_SUP[unfolded SUPR_fun_expand])
-
-lemma (in sigma_finite_measure) disjoint_sigma_finite:
-  "\<exists>A::nat\<Rightarrow>'a set. range A \<subseteq> sets M \<and> (\<Union>i. A i) = space M \<and>
-    (\<forall>i. \<mu> (A i) \<noteq> \<omega>) \<and> disjoint_family A"
-proof -
-  obtain A :: "nat \<Rightarrow> 'a set" where
-    range: "range A \<subseteq> sets M" and
-    space: "(\<Union>i. A i) = space M" and
-    measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>"
-    using sigma_finite by auto
-
-  note range' = range_disjointed_sets[OF range] range
-
-  { fix i
-    have "\<mu> (disjointed A i) \<le> \<mu> (A i)"
-      using range' disjointed_subset[of A i] by (auto intro!: measure_mono)
-    then have "\<mu> (disjointed A i) \<noteq> \<omega>"
-      using measure[of i] by auto }
-  with disjoint_family_disjointed UN_disjointed_eq[of A] space range'
-  show ?thesis by (auto intro!: exI[of _ "disjointed A"])
-qed
-
 lemma (in sigma_finite_measure) Ex_finite_integrable_function:
   shows "\<exists>h\<in>borel_measurable M. positive_integral h \<noteq> \<omega> \<and> (\<forall>x\<in>space M. 0 < h x \<and> h x < \<omega>)"
 proof -
@@ -206,7 +11,6 @@
     measure: "\<And>i. \<mu> (A i) \<noteq> \<omega>" and
     disjoint: "disjoint_family A"
     using disjoint_sigma_finite by auto
-
   let "?B i" = "2^Suc i * \<mu> (A i)"
   have "\<forall>i. \<exists>x. 0 < x \<and> x < inverse (?B i)"
   proof
@@ -225,20 +29,22 @@
   qed
   from choice[OF this] obtain n where n: "\<And>i. 0 < n i"
     "\<And>i. n i < inverse (2^Suc i * \<mu> (A i))" by auto
-
   let "?h x" = "\<Sum>\<^isub>\<infinity> i. n i * indicator (A i) x"
   show ?thesis
   proof (safe intro!: bexI[of _ ?h] del: notI)
-    have "positive_integral ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))"
-      apply (subst positive_integral_psuminf)
-      using range apply (fastsimp intro!: borel_measurable_pinfreal_times borel_measurable_const borel_measurable_indicator)
-      apply (subst positive_integral_cmult_indicator)
-      using range by auto
+    have "\<And>i. A i \<in> sets M"
+      using range by fastsimp+
+    then have "positive_integral ?h = (\<Sum>\<^isub>\<infinity> i. n i * \<mu> (A i))"
+      by (simp add: positive_integral_psuminf positive_integral_cmult_indicator)
     also have "\<dots> \<le> (\<Sum>\<^isub>\<infinity> i. Real ((1 / 2)^Suc i))"
     proof (rule psuminf_le)
       fix N show "n N * \<mu> (A N) \<le> Real ((1 / 2) ^ Suc N)"
         using measure[of N] n[of N]
-        by (cases "n N") (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff mult_le_0_iff mult_less_0_iff power_less_zero_eq power_le_zero_eq inverse_eq_divide less_divide_eq power_divide split: split_if_asm)
+        by (cases "n N")
+           (auto simp: pinfreal_noteq_omega_Ex field_simps zero_le_mult_iff
+                       mult_le_0_iff mult_less_0_iff power_less_zero_eq
+                       power_le_zero_eq inverse_eq_divide less_divide_eq
+                       power_divide split: split_if_asm)
     qed
     also have "\<dots> = Real 1"
       by (rule suminf_imp_psuminf, rule power_half_series, auto)
@@ -251,13 +57,37 @@
     then show "0 < ?h x" and "?h x < \<omega>" using n[of i] by auto
   next
     show "?h \<in> borel_measurable M" using range
-      by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times borel_measurable_indicator)
+      by (auto intro!: borel_measurable_psuminf borel_measurable_pinfreal_times)
   qed
 qed
 
 definition (in measure_space)
   "absolutely_continuous \<nu> = (\<forall>N\<in>null_sets. \<nu> N = (0 :: pinfreal))"
 
+lemma (in finite_measure_space) absolutely_continuousI:
+  assumes "finite_measure_space M \<nu>"
+  assumes v: "\<And>x. \<lbrakk> x \<in> space M ; \<mu> {x} = 0 \<rbrakk> \<Longrightarrow> \<nu> {x} = 0"
+  shows "absolutely_continuous \<nu>"
+proof (unfold absolutely_continuous_def sets_eq_Pow, safe)
+  fix N assume "\<mu> N = 0" "N \<subseteq> space M"
+  interpret v: finite_measure_space M \<nu> by fact
+  have "\<nu> N = \<nu> (\<Union>x\<in>N. {x})" by simp
+  also have "\<dots> = (\<Sum>x\<in>N. \<nu> {x})"
+  proof (rule v.measure_finitely_additive''[symmetric])
+    show "finite N" using `N \<subseteq> space M` finite_space by (auto intro: finite_subset)
+    show "disjoint_family_on (\<lambda>i. {i}) N" unfolding disjoint_family_on_def by auto
+    fix x assume "x \<in> N" thus "{x} \<in> sets M" using `N \<subseteq> space M` sets_eq_Pow by auto
+  qed
+  also have "\<dots> = 0"
+  proof (safe intro!: setsum_0')
+    fix x assume "x \<in> N"
+    hence "\<mu> {x} \<le> \<mu> N" using sets_eq_Pow `N \<subseteq> space M` by (auto intro!: measure_mono)
+    hence "\<mu> {x} = 0" using `\<mu> N = 0` by simp
+    thus "\<nu> {x} = 0" using v[of x] `x \<in> N` `N \<subseteq> space M` by auto
+  qed
+  finally show "\<nu> N = 0" .
+qed
+
 lemma (in finite_measure) Radon_Nikodym_aux_epsilon:
   fixes e :: real assumes "0 < e"
   assumes "finite_measure M s"
@@ -370,7 +200,7 @@
 
   interpret M': finite_measure M s by fact
 
-  let "?r S" = "M\<lparr> space := S, sets := (\<lambda>C. S \<inter> C)`sets M\<rparr>"
+  let "?r S" = "restricted_space S"
 
   { fix S n
     assume S: "S \<in> sets M"
@@ -838,7 +668,7 @@
       = (f x * indicator (Q i) x) * indicator A x"
       unfolding indicator_def by auto
 
-    have fm: "finite_measure (M\<lparr>space := Q i, sets := op \<inter> (Q i) ` sets M\<rparr>) \<mu>"
+    have fm: "finite_measure (restricted_space (Q i)) \<mu>"
       (is "finite_measure ?R \<mu>") by (rule restricted_finite_measure[OF Q_sets[of i]])
     then interpret R: finite_measure ?R .
     have fmv: "finite_measure ?R \<nu>"
@@ -935,47 +765,6 @@
   qed
 qed
 
-lemma (in measure_space) positive_integral_translated_density:
-  assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M"
-  shows "measure_space.positive_integral M (\<lambda>A. positive_integral (\<lambda>x. f x * indicator A x)) g =
-    positive_integral (\<lambda>x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _")
-proof -
-  from measure_space_density[OF assms(1)]
-  interpret T: measure_space M ?T .
-
-  from borel_measurable_implies_simple_function_sequence[OF assms(2)]
-  obtain G where G: "\<And>i. simple_function (G i)" "G \<up> g" by blast
-  note G_borel = borel_measurable_simple_function[OF this(1)]
-
-  from T.positive_integral_isoton[OF `G \<up> g` G_borel]
-  have *: "(\<lambda>i. T.positive_integral (G i)) \<up> T.positive_integral g" .
-
-  { fix i
-    have [simp]: "finite (G i ` space M)"
-      using G(1) unfolding simple_function_def by auto
-    have "T.positive_integral (G i) = T.simple_integral (G i)"
-      using G T.positive_integral_eq_simple_integral by simp
-    also have "\<dots> = positive_integral (\<lambda>x. f x * (\<Sum>y\<in>G i`space M. y * indicator (G i -` {y} \<inter> space M) x))"
-      apply (simp add: T.simple_integral_def)
-      apply (subst positive_integral_cmult[symmetric])
-      using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
-      apply (subst positive_integral_setsum[symmetric])
-      using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage)
-      by (simp add: setsum_right_distrib field_simps)
-    also have "\<dots> = positive_integral (\<lambda>x. f x * G i x)"
-      by (auto intro!: positive_integral_cong
-               simp: indicator_def if_distrib setsum_cases)
-    finally have "T.positive_integral (G i) = positive_integral (\<lambda>x. f x * G i x)" . }
-  with * have eq_Tg: "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> T.positive_integral g" by simp
-
-  from G(2) have "(\<lambda>i x. f x * G i x) \<up> (\<lambda>x. f x * g x)"
-    unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right)
-  then have "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> positive_integral (\<lambda>x. f x * g x)"
-    using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pinfreal_times)
-  with eq_Tg show "T.positive_integral g = positive_integral (\<lambda>x. f x * g x)"
-    unfolding isoton_def by simp
-qed
-
 lemma (in sigma_finite_measure) Radon_Nikodym:
   assumes "measure_space M \<nu>"
   assumes "absolutely_continuous \<nu>"
--- a/src/HOL/Probability/Sigma_Algebra.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/Sigma_Algebra.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -6,7 +6,7 @@
 
 header {* Sigma Algebras *}
 
-theory Sigma_Algebra imports Main Countable FuncSet begin
+theory Sigma_Algebra imports Main Countable FuncSet Indicator_Function begin
 
 text {* Sigma algebras are an elementary concept in measure
   theory. To measure --- that is to integrate --- functions, we first have
@@ -95,10 +95,13 @@
 lemma (in algebra) Int_space_eq2 [simp]: "x \<in> sets M \<Longrightarrow> x \<inter> space M = x"
   by (metis Int_absorb2 sets_into_space)
 
+section {* Restricted algebras *}
+
+abbreviation (in algebra)
+  "restricted_space A \<equiv> \<lparr> space = A, sets = (\<lambda>S. (A \<inter> S)) ` sets M \<rparr>"
+
 lemma (in algebra) restricted_algebra:
-  assumes "S \<in> sets M"
-  shows "algebra (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>)"
-    (is "algebra ?r")
+  assumes "A \<in> sets M" shows "algebra (restricted_space A)"
   using assms by unfold_locales auto
 
 subsection {* Sigma Algebras *}
@@ -107,6 +110,13 @@
   assumes countable_nat_UN [intro]:
          "!!A. range A \<subseteq> sets M \<Longrightarrow> (\<Union>i::nat. A i) \<in> sets M"
 
+lemma sigma_algebra_cong:
+  fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
+  assumes *: "sigma_algebra M"
+  and cong: "space M = space M'" "sets M = sets M'"
+  shows "sigma_algebra M'"
+using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .
+
 lemma countable_UN_eq:
   fixes A :: "'i::countable \<Rightarrow> 'a set"
   shows "(range A \<subseteq> sets M \<longrightarrow> (\<Union>i. A i) \<in> sets M) \<longleftrightarrow>
@@ -320,15 +330,14 @@
 
 lemma (in sigma_algebra) restricted_sigma_algebra:
   assumes "S \<in> sets M"
-  shows "sigma_algebra (M\<lparr> space := S, sets := (\<lambda>A. S \<inter> A) ` sets M \<rparr>)"
-    (is "sigma_algebra ?r")
+  shows "sigma_algebra (restricted_space S)"
   unfolding sigma_algebra_def sigma_algebra_axioms_def
 proof safe
-  show "algebra ?r" using restricted_algebra[OF assms] .
+  show "algebra (restricted_space S)" using restricted_algebra[OF assms] .
 next
-  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets ?r"
+  fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets (restricted_space S)"
   from restriction_in_sets[OF assms this[simplified]]
-  show "(\<Union>i. A i) \<in> sets ?r" by simp
+  show "(\<Union>i. A i) \<in> sets (restricted_space S)" by simp
 qed
 
 section {* Measurable functions *}
@@ -560,6 +569,19 @@
       (metis insert_absorb insert_subset le_SucE le_antisym not_leE)
 qed
 
+lemma setsum_indicator_disjoint_family:
+  fixes f :: "'d \<Rightarrow> 'e::semiring_1"
+  assumes d: "disjoint_family_on A P" and "x \<in> A j" and "finite P" and "j \<in> P"
+  shows "(\<Sum>i\<in>P. f i * indicator (A i) x) = f j"
+proof -
+  have "P \<inter> {i. x \<in> A i} = {j}"
+    using d `x \<in> A j` `j \<in> P` unfolding disjoint_family_on_def
+    by auto
+  thus ?thesis
+    unfolding indicator_def
+    by (simp add: if_distrib setsum_cases[OF `finite P`])
+qed
+
 definition disjointed :: "(nat \<Rightarrow> 'a set) \<Rightarrow> nat \<Rightarrow> 'a set "
   where "disjointed A n = A n - (\<Union>i\<in>{0..<n}. A i)"
 
@@ -626,6 +648,67 @@
   thus "(\<Union>i::nat. A i) \<in> sets M" by (simp add: UN_disjointed_eq)
 qed
 
+subsection {* Sigma algebra generated by function preimages *}
+
+definition (in sigma_algebra)
+  "vimage_algebra S f = \<lparr> space = S, sets = (\<lambda>A. f -` A \<inter> S) ` sets M \<rparr>"
+
+lemma (in sigma_algebra) in_vimage_algebra[simp]:
+  "A \<in> sets (vimage_algebra S f) \<longleftrightarrow> (\<exists>B\<in>sets M. A = f -` B \<inter> S)"
+  by (simp add: vimage_algebra_def image_iff)
+
+lemma (in sigma_algebra) space_vimage_algebra[simp]:
+  "space (vimage_algebra S f) = S"
+  by (simp add: vimage_algebra_def)
+
+lemma (in sigma_algebra) sigma_algebra_vimage:
+  fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M"
+  shows "sigma_algebra (vimage_algebra S f)"
+proof
+  fix A assume "A \<in> sets (vimage_algebra S f)"
+  then guess B unfolding in_vimage_algebra ..
+  then show "space (vimage_algebra S f) - A \<in> sets (vimage_algebra S f)"
+    using sets_into_space assms
+    by (auto intro!: bexI[of _ "space M - B"])
+next
+  fix A assume "A \<in> sets (vimage_algebra S f)"
+  then guess A' unfolding in_vimage_algebra .. note A' = this
+  fix B assume "B \<in> sets (vimage_algebra S f)"
+  then guess B' unfolding in_vimage_algebra .. note B' = this
+  then show "A \<union> B \<in> sets (vimage_algebra S f)"
+    using sets_into_space assms A' B'
+    by (auto intro!: bexI[of _ "A' \<union> B'"])
+next
+  fix A::"nat \<Rightarrow> 'c set" assume "range A \<subseteq> sets (vimage_algebra S f)"
+  then have "\<forall>i. \<exists>B. A i = f -` B \<inter> S \<and> B \<in> sets M"
+    by (simp add: subset_eq) blast
+  from this[THEN choice] obtain B
+    where B: "\<And>i. A i = f -` B i \<inter> S" "range B \<subseteq> sets M" by auto
+  show "(\<Union>i. A i) \<in> sets (vimage_algebra S f)"
+    using B by (auto intro!: bexI[of _ "\<Union>i. B i"])
+qed auto
+
+lemma (in sigma_algebra) measurable_vimage_algebra:
+  fixes S :: "'c set" assumes "f \<in> S \<rightarrow> space M"
+  shows "f \<in> measurable (vimage_algebra S f) M"
+    unfolding measurable_def using assms by force
+
+section {* Conditional space *}
+
+definition (in algebra)
+  "image_space X = \<lparr> space = X`space M, sets = (\<lambda>S. X`S) ` sets M \<rparr>"
+
+definition (in algebra)
+  "conditional_space X A = algebra.image_space (restricted_space A) X"
+
+lemma (in algebra) space_conditional_space:
+  assumes "A \<in> sets M" shows "space (conditional_space X A) = X`A"
+proof -
+  interpret r: algebra "restricted_space A" using assms by (rule restricted_algebra)
+  show ?thesis unfolding conditional_space_def r.image_space_def
+    by simp
+qed
+
 subsection {* A Two-Element Series *}
 
 definition binaryset :: "'a set \<Rightarrow> 'a set \<Rightarrow> nat \<Rightarrow> 'a set "
--- a/src/HOL/Probability/ex/Dining_Cryptographers.thy	Thu Sep 02 18:45:23 2010 +0200
+++ b/src/HOL/Probability/ex/Dining_Cryptographers.thy	Thu Sep 02 20:44:33 2010 +0200
@@ -26,14 +26,13 @@
   let ?measure = "\<lambda>s::'a set. real (card s) / real (card S)"
 
   show "finite_measure_space M \<mu>"
-  proof (rule finite_Pow_additivity_sufficient, simp_all)
-    show "positive \<mu>" by (simp add: positive_def)
-
-    show "additive M \<mu>"
-      by (simp add: additive_def inverse_eq_divide field_simps Real_real
+  proof (rule finite_measure_spaceI)
+    fix A B :: "'a set" assume "A \<inter> B = {}" "A \<subseteq> space M" "B \<subseteq> space M"
+    then show "\<mu> (A \<union> B) = \<mu> A + \<mu> B"
+      by (simp add: inverse_eq_divide field_simps Real_real
                     divide_le_0_iff zero_le_divide_iff
                     card_Un_disjoint finite_subset[OF _ finite])
-  qed
+  qed auto
 qed simp_all
 
 lemma set_of_list_extend: