cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
authorhoelzl
Wed, 07 Oct 2015 17:11:16 +0200
changeset 61359 e985b52c3eb3
parent 61358 131fc8c10402
child 61360 a273bdac0934
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
src/HOL/Library/FuncSet.thy
src/HOL/Probability/Distributions.thy
src/HOL/Probability/Finite_Product_Measure.thy
src/HOL/Probability/Giry_Monad.thy
src/HOL/Probability/Independent_Family.thy
src/HOL/Probability/Infinite_Product_Measure.thy
src/HOL/Probability/Measure_Space.thy
src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy
src/HOL/Probability/Probability.thy
src/HOL/Probability/Probability_Measure.thy
src/HOL/Probability/Projective_Family.thy
src/HOL/Probability/Projective_Limit.thy
--- a/src/HOL/Library/FuncSet.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Library/FuncSet.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -184,6 +184,9 @@
 
 subsection \<open>Bounded Abstraction: @{term restrict}\<close>
 
+lemma restrict_cong: "I = J \<Longrightarrow> (\<And>i. i \<in> J =simp=> f i = g i) \<Longrightarrow> restrict f I = restrict g J"
+  by (auto simp: restrict_def fun_eq_iff simp_implies_def)
+
 lemma restrict_in_funcset: "(\<And>x. x \<in> A \<Longrightarrow> f x \<in> B) \<Longrightarrow> (\<lambda>x\<in>A. f x) \<in> A \<rightarrow> B"
   by (simp add: Pi_def restrict_def)
 
--- a/src/HOL/Probability/Distributions.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Distributions.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -973,11 +973,7 @@
         by (intro integral_cong) auto
       finally have "Suc k * (\<integral>x. indicator {0..b} x *\<^sub>R ?M k x \<partial>lborel) =
         exp (- b\<^sup>2) * b ^ (Suc k) + 2 * (\<integral>x. indicator {0..b} x *\<^sub>R ?M (k + 2) x \<partial>lborel)"
-        apply (simp del: real_scaleR_def integral_mult_right add: integral_mult_right[symmetric])
-        apply (subst integral_mult_right_zero[symmetric])
-        apply (intro integral_cong)
-        apply simp_all
-        done
+        by (simp del: real_scaleR_def integral_mult_right add: integral_mult_right[symmetric])
       then show "(k + 1) / 2 * (\<integral>x. indicator {..b} x *\<^sub>R (indicator {0..} x *\<^sub>R ?M k x)\<partial>lborel) - exp (- b\<^sup>2) * b ^ (k + 1) / 2 = ?f b"
         by (simp add: field_simps atLeastAtMost_def indicator_inter_arith)
     qed
--- a/src/HOL/Probability/Finite_Product_Measure.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Finite_Product_Measure.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -351,6 +351,15 @@
 lemma space_PiM: "space (\<Pi>\<^sub>M i\<in>I. M i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
   using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro space_extend_measure) simp
 
+lemma prod_emb_subset_PiM[simp]: "prod_emb I M K X \<subseteq> space (PiM I M)"
+  by (auto simp: prod_emb_def space_PiM)
+
+lemma space_PiM_empty_iff[simp]: "space (PiM I M) = {} \<longleftrightarrow>  (\<exists>i\<in>I. space (M i) = {})"
+  by (auto simp: space_PiM PiE_eq_empty_iff)
+
+lemma undefined_in_PiM_empty[simp]: "(\<lambda>x. undefined) \<in> space (PiM {} M)"
+  by (auto simp: space_PiM)
+
 lemma sets_PiM: "sets (\<Pi>\<^sub>M i\<in>I. M i) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
   using prod_algebra_sets_into_space unfolding PiM_def prod_algebra_def by (intro sets_extend_measure) simp
 
@@ -621,6 +630,20 @@
   finally show "{\<omega> \<in> space ?P. case_prod (\<lambda>f. fun_upd f i) \<omega> j \<in> A} \<in> sets ?P" .
 qed (auto simp: space_pair_measure space_PiM PiE_def)
 
+lemma measurable_fun_upd:
+  assumes I: "I = J \<union> {i}"
+  assumes f[measurable]: "f \<in> measurable N (PiM J M)"
+  assumes h[measurable]: "h \<in> measurable N (M i)"
+  shows "(\<lambda>x. (f x) (i := h x)) \<in> measurable N (PiM I M)"
+proof (intro measurable_PiM_single')
+  fix j assume "j \<in> I" then show "(\<lambda>\<omega>. ((f \<omega>)(i := h \<omega>)) j) \<in> measurable N (M j)"
+    unfolding I by (cases "j = i") auto
+next
+  show "(\<lambda>x. (f x)(i := h x)) \<in> space N \<rightarrow> (\<Pi>\<^sub>E i\<in>I. space (M i))"
+    using I f[THEN measurable_space] h[THEN measurable_space]
+    by (auto simp: space_PiM PiE_iff extensional_def)
+qed
+
 lemma measurable_component_update:
   "x \<in> space (Pi\<^sub>M I M) \<Longrightarrow> i \<notin> I \<Longrightarrow> (\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^sub>M (insert i I) M)"
   by simp
@@ -673,6 +696,25 @@
   unfolding prod_emb_def space_PiM[symmetric]
   by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton)
 
+lemma merge_in_prod_emb:
+  assumes "y \<in> space (PiM I M)" "x \<in> X" and X: "X \<in> sets (Pi\<^sub>M J M)" and "J \<subseteq> I"
+  shows "merge J I (x, y) \<in> prod_emb I M J X"
+  using assms sets.sets_into_space[OF X]
+  by (simp add: merge_def prod_emb_def subset_eq space_PiM PiE_def extensional_restrict Pi_iff
+           cong: if_cong restrict_cong)
+     (simp add: extensional_def)
+
+lemma prod_emb_eq_emptyD:
+  assumes J: "J \<subseteq> I" and ne: "space (PiM I M) \<noteq> {}" and X: "X \<in> sets (Pi\<^sub>M J M)"
+    and *: "prod_emb I M J X = {}"
+  shows "X = {}"
+proof safe
+  fix x assume "x \<in> X"
+  obtain \<omega> where "\<omega> \<in> space (PiM I M)"
+    using ne by blast
+  from merge_in_prod_emb[OF this \<open>x\<in>X\<close> X J] * show "x \<in> {}" by auto 
+qed
+
 lemma sets_in_Pi_aux:
   "finite I \<Longrightarrow> (\<And>j. j \<in> I \<Longrightarrow> {x\<in>space (M j). x \<in> F j} \<in> sets (M j)) \<Longrightarrow>
   {x\<in>space (PiM I M). x \<in> Pi I F} \<in> sets (PiM I M)"
@@ -817,6 +859,36 @@
     by (subst (asm) prod_emb_PiE_same_index) (auto intro!: sets.sets_into_space)
 qed simp
 
+lemma (in product_sigma_finite) PiM_eqI:
+  assumes "finite I" "sets P = PiM I M"
+  assumes eq: "\<And>A. (\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> P (Pi\<^sub>E I A) = (\<Prod>i\<in>I. emeasure (M i) (A i))"
+  shows "P = PiM I M"
+proof -
+  interpret finite_product_sigma_finite M I
+    proof qed fact
+  from sigma_finite_pairs guess C .. note C = this
+  show ?thesis
+  proof (rule measure_eqI_generator_eq[symmetric, OF Int_stable_prod_algebra prod_algebra_sets_into_space])
+    show "sets (PiM I M) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
+      by (rule sets_PiM)
+    then show "sets P = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) (prod_algebra I M)"
+      unfolding `sets P = sets (PiM I M)` by simp
+    def A \<equiv> "\<lambda>n. \<Pi>\<^sub>E i\<in>I. C i n"
+    show "range A \<subseteq> prod_algebra I M" "\<And>i. emeasure (Pi\<^sub>M I M) (A i) \<noteq> \<infinity>"
+      using C \<open>finite I\<close>
+      by (auto intro!: prod_algebraI_finite simp: A_def emeasure_PiM subset_eq setprod_PInf emeasure_nonneg)
+    show "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
+      using C by (simp add: A_def space_PiM)
+
+    fix X assume X: "X \<in> prod_algebra I M"
+    then obtain J E where X: "X = prod_emb I M J (PIE j:J. E j)"
+      and J: "finite J" "J \<subseteq> I" "\<And>j. j \<in> J \<Longrightarrow> E j \<in> sets (M j)"
+      by (force elim!: prod_algebraE)
+    show "emeasure (PiM I M) X = emeasure P X"
+      unfolding X by (subst (1 2) prod_emb_Pi) (auto simp: eq emeasure_PiM J \<open>finite I\<close>)
+  qed
+qed
+
 lemma (in product_sigma_finite) sigma_finite: 
   assumes "finite I"
   shows "sigma_finite_measure (PiM I M)"
@@ -843,84 +915,24 @@
   using emeasure_PiM[OF finite_index] by auto
 
 lemma (in product_sigma_finite) nn_integral_empty:
-  assumes pos: "0 \<le> f (\<lambda>k. undefined)"
-  shows "integral\<^sup>N (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)"
-proof -
-  interpret finite_product_sigma_finite M "{}" by standard (fact finite.emptyI)
-  have "\<And>A. emeasure (Pi\<^sub>M {} M) (Pi\<^sub>E {} A) = 1"
-    using assms by (subst measure_times) auto
-  then show ?thesis
-    unfolding nn_integral_def simple_function_def simple_integral_def[abs_def]
-  proof (simp add: space_PiM_empty sets_PiM_empty, intro antisym)
-    show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined))"
-      by (intro SUP_upper) (auto simp: le_fun_def split: split_max)
-    show "(SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)" using pos
-      by (intro SUP_least) (auto simp: le_fun_def simp: max_def split: split_if_asm)
-  qed
-qed
+  "0 \<le> f (\<lambda>k. undefined) \<Longrightarrow> integral\<^sup>N (Pi\<^sub>M {} M) f = f (\<lambda>k. undefined)"
+  by (simp add: PiM_empty nn_integral_count_space_finite max.absorb2)
 
 lemma (in product_sigma_finite) distr_merge:
   assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J"
   shows "distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J) = Pi\<^sub>M (I \<union> J) M"
    (is "?D = ?P")
-proof -
+proof (rule PiM_eqI)
   interpret I: finite_product_sigma_finite M I by standard fact
   interpret J: finite_product_sigma_finite M J by standard fact
-  have "finite (I \<union> J)" using fin by auto
-  interpret IJ: finite_product_sigma_finite M "I \<union> J" by standard fact
-  interpret P: pair_sigma_finite "Pi\<^sub>M I M" "Pi\<^sub>M J M" by standard
-  let ?g = "merge I J"
-
-  from IJ.sigma_finite_pairs obtain F where
-    F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)"
-       "incseq (\<lambda>k. \<Pi>\<^sub>E i\<in>I \<union> J. F i k)"
-       "(\<Union>k. \<Pi>\<^sub>E i\<in>I \<union> J. F i k) = space ?P"
-       "\<And>k. \<forall>i\<in>I\<union>J. emeasure (M i) (F i k) \<noteq> \<infinity>"
-    by auto
-  let ?F = "\<lambda>k. \<Pi>\<^sub>E i\<in>I \<union> J. F i k"
-  
-  show ?thesis
-  proof (rule measure_eqI_generator_eq[symmetric])
-    show "Int_stable (prod_algebra (I \<union> J) M)"
-      by (rule Int_stable_prod_algebra)
-    show "prod_algebra (I \<union> J) M \<subseteq> Pow (\<Pi>\<^sub>E i \<in> I \<union> J. space (M i))"
-      by (rule prod_algebra_sets_into_space)
-    show "sets ?P = sigma_sets (\<Pi>\<^sub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
-      by (rule sets_PiM)
-    then show "sets ?D = sigma_sets (\<Pi>\<^sub>E i\<in>I \<union> J. space (M i)) (prod_algebra (I \<union> J) M)"
-      by simp
-
-    show "range ?F \<subseteq> prod_algebra (I \<union> J) M" using F
-      using fin by (auto simp: prod_algebra_eq_finite)
-    show "(\<Union>i. \<Pi>\<^sub>E ia\<in>I \<union> J. F ia i) = (\<Pi>\<^sub>E i\<in>I \<union> J. space (M i))"
-      using F(3) by (simp add: space_PiM)
-  next
-    fix k
-    from F `finite I` setprod_PInf[of "I \<union> J", OF emeasure_nonneg, of M]
-    show "emeasure ?P (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto
-  next
-    fix A assume A: "A \<in> prod_algebra (I \<union> J) M"
-    with fin obtain F where A_eq: "A = (Pi\<^sub>E (I \<union> J) F)" and F: "\<forall>i\<in>J. F i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
-      by (auto simp add: prod_algebra_eq_finite)
-    let ?B = "Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M"
-    let ?X = "?g -` A \<inter> space ?B"
-    have "Pi\<^sub>E I F \<subseteq> space (Pi\<^sub>M I M)" "Pi\<^sub>E J F \<subseteq> space (Pi\<^sub>M J M)"
-      using F[rule_format, THEN sets.sets_into_space] by (force simp: space_PiM)+
-    then have X: "?X = (Pi\<^sub>E I F \<times> Pi\<^sub>E J F)"
-      unfolding A_eq by (subst merge_vimage) (auto simp: space_pair_measure space_PiM)
-    have "emeasure ?D A = emeasure ?B ?X"
-      using A by (intro emeasure_distr measurable_merge) (auto simp: sets_PiM)
-    also have "emeasure ?B ?X = (\<Prod>i\<in>I. emeasure (M i) (F i)) * (\<Prod>i\<in>J. emeasure (M i) (F i))"
-      using `finite J` `finite I` F unfolding X
-      by (simp add: J.emeasure_pair_measure_Times I.measure_times J.measure_times)
-    also have "\<dots> = (\<Prod>i\<in>I \<union> J. emeasure (M i) (F i))"
-      using `finite J` `finite I` `I \<inter> J = {}`  by (simp add: setprod.union_inter_neutral)
-    also have "\<dots> = emeasure ?P (Pi\<^sub>E (I \<union> J) F)"
-      using `finite J` `finite I` F unfolding A
-      by (intro IJ.measure_times[symmetric]) auto
-    finally show "emeasure ?P A = emeasure ?D A" using A_eq by simp
-  qed
-qed
+  fix A assume A: "\<And>i. i \<in> I \<union> J \<Longrightarrow> A i \<in> sets (M i)"
+  have *: "(merge I J -` Pi\<^sub>E (I \<union> J) A \<inter> space (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M)) = PiE I A \<times> PiE J A"
+    using A[THEN sets.sets_into_space] by (auto simp: space_PiM space_pair_measure)
+  from A fin show "emeasure (distr (Pi\<^sub>M I M \<Otimes>\<^sub>M Pi\<^sub>M J M) (Pi\<^sub>M (I \<union> J) M) (merge I J)) (Pi\<^sub>E (I \<union> J) A) =
+      (\<Prod>i\<in>I \<union> J. emeasure (M i) (A i))"
+    by (subst emeasure_distr)
+       (auto simp: * J.emeasure_pair_measure_Times I.measure_times J.measure_times setprod.union_disjoint)
+qed (insert fin, simp_all)
 
 lemma (in product_sigma_finite) product_nn_integral_fold:
   assumes IJ: "I \<inter> J = {}" "finite I" "finite J"
@@ -1043,23 +1055,13 @@
 
 lemma (in product_sigma_finite) distr_component:
   "distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x) = Pi\<^sub>M {i} M" (is "?D = ?P")
-proof (intro measure_eqI[symmetric])
-  interpret I: finite_product_sigma_finite M "{i}" by standard simp
-
-  have eq: "\<And>x. x \<in> extensional {i} \<Longrightarrow> (\<lambda>j\<in>{i}. x i) = x"
-    by (auto simp: extensional_def restrict_def)
-
-  have [measurable]: "\<And>j. j \<in> {i} \<Longrightarrow> (\<lambda>x. x) \<in> measurable (M i) (M j)" by simp
-
-  fix A assume A: "A \<in> sets ?P"
-  then have "emeasure ?P A = (\<integral>\<^sup>+x. indicator A x \<partial>?P)" 
-    by simp
-  also have "\<dots> = (\<integral>\<^sup>+x. indicator ((\<lambda>x. \<lambda>i\<in>{i}. x) -` A \<inter> space (M i)) (x i) \<partial>PiM {i} M)" 
-    by (intro nn_integral_cong) (auto simp: space_PiM indicator_def PiE_def eq)
-  also have "\<dots> = emeasure ?D A"
-    using A by (simp add: product_nn_integral_singleton emeasure_distr)
-  finally show "emeasure (Pi\<^sub>M {i} M) A = emeasure ?D A" .
-qed simp
+proof (intro PiM_eqI)
+  fix A assume "\<And>ia. ia \<in> {i} \<Longrightarrow> A ia \<in> sets (M ia)"
+  moreover then have "(\<lambda>x. \<lambda>i\<in>{i}. x) -` Pi\<^sub>E {i} A \<inter> space (M i) = A i"
+    by (auto dest: sets.sets_into_space)
+  ultimately show "emeasure (distr (M i) (Pi\<^sub>M {i} M) (\<lambda>x. \<lambda>i\<in>{i}. x)) (Pi\<^sub>E {i} A) = (\<Prod>i\<in>{i}. emeasure (M i) (A i))"
+    by (subst emeasure_distr) (auto intro!: sets_PiM_I_finite measurable_restrict)
+qed simp_all
 
 lemma (in product_sigma_finite)
   assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^sub>M (I \<union> J) M)"
--- a/src/HOL/Probability/Giry_Monad.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Giry_Monad.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -1405,6 +1405,10 @@
   apply (auto intro!: return_measurable simp: distr_distr[symmetric] join_return')
   done
 
+lemma bind_return_distr':
+  "space M \<noteq> {} \<Longrightarrow> f \<in> measurable M N \<Longrightarrow> bind M (\<lambda>x. return N (f x)) = distr M N f"
+  using bind_return_distr[of M f N] by (simp add: comp_def)
+
 lemma bind_assoc:
   fixes f :: "'a \<Rightarrow> 'b measure" and g :: "'b \<Rightarrow> 'c measure"
   assumes M1: "f \<in> measurable M (subprob_algebra N)" and M2: "g \<in> measurable N (subprob_algebra R)"
--- a/src/HOL/Probability/Independent_Family.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Independent_Family.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -1263,7 +1263,7 @@
   note rv_Y = this[measurable]
     
   interpret Y: prob_space "distr M borel (Y i)" for i
-    using I(2) by (cases "i \<in> I") (auto intro!: prob_space_distr simp: Y_def indep_vars_def)
+    using I(2) by (cases "i \<in> I") (auto intro!: prob_space_distr simp: indep_vars_def prob_space_return)
   interpret product_sigma_finite "\<lambda>i. distr M borel (Y i)"
     ..
   
@@ -1305,7 +1305,7 @@
   note int_Y = this
     
   interpret Y: prob_space "distr M borel (Y i)" for i
-    using I(2) by (cases "i \<in> I") (auto intro!: prob_space_distr simp: Y_def indep_vars_def)
+    using I(2) by (cases "i \<in> I") (auto intro!: prob_space_distr simp: indep_vars_def prob_space_return)
   interpret product_sigma_finite "\<lambda>i. distr M borel (Y i)"
     ..
   
--- a/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Infinite_Product_Measure.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -8,341 +8,55 @@
   imports Probability_Measure Caratheodory Projective_Family
 begin
 
-lemma (in product_prob_space) emeasure_PiM_emb_not_empty:
-  assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. X i \<in> sets (M i)"
-  shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)"
-proof cases
-  assume "finite I" with X show ?thesis by simp
-next
-  let ?\<Omega> = "\<Pi>\<^sub>E i\<in>I. space (M i)"
-  let ?G = generator
-  assume "\<not> finite I"
-  then have I_not_empty: "I \<noteq> {}" by auto
-  interpret G!: algebra ?\<Omega> generator by (rule algebra_generator) fact
-  note mu_G_mono =
-    G.additive_increasing[OF positive_mu_G[OF I_not_empty] additive_mu_G[OF I_not_empty],
-      THEN increasingD]
-  write mu_G  ("\<mu>G")
-
-  { fix Z J assume J: "J \<noteq> {}" "finite J" "J \<subseteq> I" and Z: "Z \<in> ?G"
-
-    from `infinite I` `finite J` obtain k where k: "k \<in> I" "k \<notin> J"
-      by (metis rev_finite_subset subsetI)
-    moreover from Z guess K' X' by (rule generatorE)
-    moreover def K \<equiv> "insert k K'"
-    moreover def X \<equiv> "emb K K' X'"
-    ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^sub>M K M)" "Z = emb I K X"
-      "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^sub>M K M) X"
-      by (auto simp: subset_insertI)
-    let ?M = "\<lambda>y. (\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^sub>M (K - J) M)"
-    { fix y assume y: "y \<in> space (Pi\<^sub>M J M)"
-      note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X]
-      moreover
-      have **: "?M y \<in> sets (Pi\<^sub>M (K - J) M)"
-        using J K y by (intro merge_sets) auto
-      ultimately
-      have ***: "((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)) \<in> ?G"
-        using J K by (intro generatorI) auto
-      have "\<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^sub>M I M)) = emeasure (Pi\<^sub>M (K - J) M) (?M y)"
-        unfolding * using K J by (subst mu_G_eq[OF _ _ _ **]) auto
-      note * ** *** this }
-    note merge_in_G = this
-
-    have "finite (K - J)" using K by auto
-
-    interpret J: finite_product_prob_space M J by standard fact+
-    interpret KmJ: finite_product_prob_space M "K - J" by standard fact+
-
-    have "\<mu>G Z = emeasure (Pi\<^sub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)"
-      using K J by simp
-    also have "\<dots> = (\<integral>\<^sup>+ x. emeasure (Pi\<^sub>M (K - J) M) (?M x) \<partial>Pi\<^sub>M J M)"
-      using K J by (subst emeasure_fold_integral) auto
-    also have "\<dots> = (\<integral>\<^sup>+ y. \<mu>G ((\<lambda>x. merge J (I - J) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)) \<partial>Pi\<^sub>M J M)"
-      (is "_ = (\<integral>\<^sup>+x. \<mu>G (?MZ x) \<partial>Pi\<^sub>M J M)")
-    proof (intro nn_integral_cong)
-      fix x assume x: "x \<in> space (Pi\<^sub>M J M)"
-      with K merge_in_G(2)[OF this]
-      show "emeasure (Pi\<^sub>M (K - J) M) (?M x) = \<mu>G (?MZ x)"
-        unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst mu_G_eq) auto
-    qed
-    finally have fold: "\<mu>G Z = (\<integral>\<^sup>+x. \<mu>G (?MZ x) \<partial>Pi\<^sub>M J M)" .
-
-    { fix x assume x: "x \<in> space (Pi\<^sub>M J M)"
-      then have "\<mu>G (?MZ x) \<le> 1"
-        unfolding merge_in_G(4)[OF x] `Z = emb I K X`
-        by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) }
-    note le_1 = this
-
-    let ?q = "\<lambda>y. \<mu>G ((\<lambda>x. merge J (I - J) (y,x)) -` Z \<inter> space (Pi\<^sub>M I M))"
-    have "?q \<in> borel_measurable (Pi\<^sub>M J M)"
-      unfolding `Z = emb I K X` using J K merge_in_G(3)
-      by (simp add: merge_in_G  mu_G_eq emeasure_fold_measurable cong: measurable_cong)
-    note this fold le_1 merge_in_G(3) }
-  note fold = this
-
-  have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>"
-  proof (rule G.caratheodory_empty_continuous[OF positive_mu_G additive_mu_G])
-    fix A assume "A \<in> ?G"
-    with generatorE guess J X . note JX = this
-    interpret JK: finite_product_prob_space M J by standard fact+ 
-    from JX show "\<mu>G A \<noteq> \<infinity>" by simp
-  next
-    fix A assume A: "range A \<subseteq> ?G" "decseq A" "(\<Inter>i. A i) = {}"
-    then have "decseq (\<lambda>i. \<mu>G (A i))"
-      by (auto intro!: mu_G_mono simp: decseq_def)
-    moreover
-    have "(INF i. \<mu>G (A i)) = 0"
-    proof (rule ccontr)
-      assume "(INF i. \<mu>G (A i)) \<noteq> 0" (is "?a \<noteq> 0")
-      moreover have "0 \<le> ?a"
-        using A positive_mu_G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def)
-      ultimately have "0 < ?a" by auto
-
-      have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^sub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (limP J M (\<lambda>J. (Pi\<^sub>M J M))) X"
-        using A by (intro allI generator_Ex) auto
-      then obtain J' X' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I" "\<And>n. X' n \<in> sets (Pi\<^sub>M (J' n) M)"
-        and A': "\<And>n. A n = emb I (J' n) (X' n)"
-        unfolding choice_iff by blast
-      moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"
-      moreover def X \<equiv> "\<lambda>n. emb (J n) (J' n) (X' n)"
-      ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. X n \<in> sets (Pi\<^sub>M (J n) M)"
-        by auto
-      with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> ?G"
-        unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto)
-
-      have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
-        unfolding J_def by force
-
-      interpret J: finite_product_prob_space M "J i" for i by standard fact+
-
-      have a_le_1: "?a \<le> 1"
-        using mu_G_spec[of "J 0" "A 0" "X 0"] J A_eq
-        by (auto intro!: INF_lower2[of 0] J.measure_le_1)
-
-      let ?M = "\<lambda>K Z y. (\<lambda>x. merge K (I - K) (y, x)) -` Z \<inter> space (Pi\<^sub>M I M)"
-
-      { fix Z k assume Z: "range Z \<subseteq> ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)"
-        then have Z_sets: "\<And>n. Z n \<in> ?G" by auto
-        fix J' assume J': "J' \<noteq> {}" "finite J'" "J' \<subseteq> I"
-        interpret J': finite_product_prob_space M J' by standard fact+
-
-        let ?q = "\<lambda>n y. \<mu>G (?M J' (Z n) y)"
-        let ?Q = "\<lambda>n. ?q n -` {?a / 2^(k+1) ..} \<inter> space (Pi\<^sub>M J' M)"
-        { fix n
-          have "?q n \<in> borel_measurable (Pi\<^sub>M J' M)"
-            using Z J' by (intro fold(1)) auto
-          then have "?Q n \<in> sets (Pi\<^sub>M J' M)"
-            by (rule measurable_sets) auto }
-        note Q_sets = this
+lemma (in product_prob_space) distr_PiM_restrict_finite:
+  assumes "finite J" "J \<subseteq> I"
+  shows "distr (PiM I M) (PiM J M) (\<lambda>x. restrict x J) = PiM J M"
+proof (rule PiM_eqI)
+  fix X assume X: "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
+  { fix J X assume J: "J \<noteq> {} \<or> I = {}" "finite J" "J \<subseteq> I" and X: "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
+    have "emeasure (PiM I M) (emb I J (PiE J X)) = (\<Prod>i\<in>J. M i (X i))"
+    proof (subst emeasure_extend_measure_Pair[OF PiM_def, where \<mu>'=lim], goal_cases)
+      case 1 then show ?case
+        by (simp add: M.emeasure_space_1 emeasure_PiM Pi_iff sets_PiM_I_finite emeasure_lim_emb)
+    next
+      case (2 J X)
+      then have "emb I J (Pi\<^sub>E J X) \<in> sets (PiM I M)"
+        by (intro measurable_prod_emb sets_PiM_I_finite) auto
+      from this[THEN sets.sets_into_space] show ?case
+        by (simp add: space_PiM)
+    qed (insert assms J X, simp_all del: sets_lim
+      add: M.emeasure_space_1 sets_lim[symmetric] emeasure_countably_additive emeasure_positive) }
+  note * = this
 
-        have "?a / 2^(k+1) \<le> (INF n. emeasure (Pi\<^sub>M J' M) (?Q n))"
-        proof (intro INF_greatest)
-          fix n
-          have "?a / 2^k \<le> \<mu>G (Z n)" using Z by auto
-          also have "\<dots> \<le> (\<integral>\<^sup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^sub>M J' M)"
-            unfolding fold(2)[OF J' `Z n \<in> ?G`]
-          proof (intro nn_integral_mono)
-            fix x assume x: "x \<in> space (Pi\<^sub>M J' M)"
-            then have "?q n x \<le> 1 + 0"
-              using J' Z fold(3) Z_sets by auto
-            also have "\<dots> \<le> 1 + ?a / 2^(k+1)"
-              using `0 < ?a` by (intro add_mono) auto
-            finally have "?q n x \<le> 1 + ?a / 2^(k+1)" .
-            with x show "?q n x \<le> indicator (?Q n) x + ?a / 2^(k+1)"
-              by (auto split: split_indicator simp del: power_Suc)
-          qed
-          also have "\<dots> = emeasure (Pi\<^sub>M J' M) (?Q n) + ?a / 2^(k+1)"
-            using `0 \<le> ?a` Q_sets J'.emeasure_space_1
-            by (subst nn_integral_add) auto
-          finally show "?a / 2^(k+1) \<le> emeasure (Pi\<^sub>M J' M) (?Q n)" using `?a \<le> 1`
-            by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^sub>M J' M) (?Q n)"])
-               (auto simp: field_simps)
-        qed
-        also have "\<dots> = emeasure (Pi\<^sub>M J' M) (\<Inter>n. ?Q n)"
-        proof (intro INF_emeasure_decseq)
-          show "range ?Q \<subseteq> sets (Pi\<^sub>M J' M)" using Q_sets by auto
-          show "decseq ?Q"
-            unfolding decseq_def
-          proof (safe intro!: vimageI[OF refl])
-            fix m n :: nat assume "m \<le> n"
-            fix x assume x: "x \<in> space (Pi\<^sub>M J' M)"
-            assume "?a / 2^(k+1) \<le> ?q n x"
-            also have "?q n x \<le> ?q m x"
-            proof (rule mu_G_mono)
-              from fold(4)[OF J', OF Z_sets x]
-              show "?M J' (Z n) x \<in> ?G" "?M J' (Z m) x \<in> ?G" by auto
-              show "?M J' (Z n) x \<subseteq> ?M J' (Z m) x"
-                using `decseq Z`[THEN decseqD, OF `m \<le> n`] by auto
-            qed
-            finally show "?a / 2^(k+1) \<le> ?q m x" .
-          qed
-        qed simp
-        finally have "(\<Inter>n. ?Q n) \<noteq> {}"
-          using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
-        then have "\<exists>w\<in>space (Pi\<^sub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto }
-      note Ex_w = this
-
-      let ?q = "\<lambda>k n y. \<mu>G (?M (J k) (A n) y)"
+  have "emeasure (PiM I M) (emb I J (PiE J X)) = (\<Prod>i\<in>J. M i (X i))"
+  proof cases
+    assume "\<not> (J \<noteq> {} \<or> I = {})"
+    then obtain i where "J = {}" "i \<in> I" by auto
+    moreover then have "emb I {} {\<lambda>x. undefined} = emb I {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
+      by (auto simp: space_PiM prod_emb_def)
+    ultimately show ?thesis
+      by (simp add: * M.emeasure_space_1)
+  qed (simp add: *[OF _ assms X])
+  with assms show "emeasure (distr (Pi\<^sub>M I M) (Pi\<^sub>M J M) (\<lambda>x. restrict x J)) (Pi\<^sub>E J X) = (\<Prod>i\<in>J. emeasure (M i) (X i))"
+    by (subst emeasure_distr_restrict[OF _ refl]) (auto intro!: sets_PiM_I_finite X)
+qed (insert assms, auto)
 
-      let ?P = "\<lambda>w k. w \<in> space (Pi\<^sub>M (J k) M) \<and> (\<forall>n. ?a / 2 ^ (Suc k) \<le> ?q k n w)"
-      have "\<exists>w. \<forall>k. ?P (w k) k \<and> restrict (w (Suc k)) (J k) = w k"
-      proof (rule dependent_nat_choice)
-        have "\<forall>n. ?a / 2 ^ 0 \<le> \<mu>G (A n)" by (auto intro: INF_lower)
-        from Ex_w[OF A(1,2) this J(1-3), of 0] show "\<exists>w. ?P w 0" by auto
-      next
-        fix w k assume Suc: "?P w k"
-        show "\<exists>w'. ?P w' (Suc k) \<and> restrict w' (J k) = w"
-        proof cases
-          assume [simp]: "J k = J (Suc k)"
-          have "?a / 2 ^ (Suc (Suc k)) \<le> ?a / 2 ^ (k + 1)"
-            using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: field_simps)
-          with Suc show ?thesis
-            by (auto intro!: exI[of _ w] simp: extensional_restrict space_PiM intro: order_trans)
-        next
-          assume "J k \<noteq> J (Suc k)"
-          with J_mono[of k "Suc k"] have "J (Suc k) - J k \<noteq> {}" (is "?D \<noteq> {}") by auto
-          have "range (\<lambda>n. ?M (J k) (A n) w) \<subseteq> ?G" "decseq (\<lambda>n. ?M (J k) (A n) w)"
-            "\<forall>n. ?a / 2 ^ (k + 1) \<le> \<mu>G (?M (J k) (A n) w)"
-            using `decseq A` fold(4)[OF J(1-3) A_eq(2), of w k] Suc
-            by (auto simp: decseq_def)
-          from Ex_w[OF this `?D \<noteq> {}`] J[of "Suc k"]
-          obtain w' where w': "w' \<in> space (Pi\<^sub>M ?D M)"
-            "\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> \<mu>G (?M ?D (?M (J k) (A n) w) w')" by auto
-          let ?w = "merge (J k) ?D (w, w')"
-          have [simp]: "\<And>x. merge (J k) (I - J k) (w, merge ?D (I - ?D) (w', x)) =
-            merge (J (Suc k)) (I - (J (Suc k))) (?w, x)"
-            using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"]
-            by (auto intro!: ext split: split_merge)
-          have *: "\<And>n. ?M ?D (?M (J k) (A n) w) w' = ?M (J (Suc k)) (A n) ?w"
-            using w'(1) J(3)[of "Suc k"]
-            by (auto simp: space_PiM split: split_merge intro!: extensional_merge_sub) force+
-          show ?thesis
-            using Suc w' J_mono[of k "Suc k"] unfolding *
-            by (intro exI[of _ ?w])
-               (auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM PiE_iff)
-        qed
-      qed
-      then obtain w where w:
-        "\<And>k. w k \<in> space (Pi\<^sub>M (J k) M)"
-        "\<And>k n. ?a / 2 ^ (Suc k) \<le> ?q k n (w k)"
-        "\<And>k. restrict (w (Suc k)) (J k) = w k"
-        by metis
-
-      { fix k
-        from w have "?a / 2 ^ (k + 1) \<le> ?q k k (w k)" by auto
-        then have "?M (J k) (A k) (w k) \<noteq> {}"
-          using positive_mu_G[OF I_not_empty, unfolded positive_def] `0 < ?a` `?a \<le> 1`
-          by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
-        then obtain x where "x \<in> ?M (J k) (A k) (w k)" by auto
-        then have "merge (J k) (I - J k) (w k, x) \<in> A k" by auto
-        then have "\<exists>x\<in>A k. restrict x (J k) = w k"
-          using `w k \<in> space (Pi\<^sub>M (J k) M)`
-          by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM) }
-      note w_ext = this
+lemma (in product_prob_space) emeasure_PiM_emb':
+  "J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> X \<in> sets (PiM J M) \<Longrightarrow> emeasure (Pi\<^sub>M I M) (emb I J X) = PiM J M X"
+  by (subst distr_PiM_restrict_finite[symmetric, of J])
+     (auto intro!: emeasure_distr_restrict[symmetric])
 
-      { fix k l i assume "k \<le> l" "i \<in> J k"
-        { fix l have "w k i = w (k + l) i"
-          proof (induct l)
-            case (Suc l)
-            from `i \<in> J k` J_mono[of k "k + l"] have "i \<in> J (k + l)" by auto
-            with w(3)[of "k + l"]
-            have "w (k + l) i = w (k + Suc l) i"
-              by (auto simp: restrict_def fun_eq_iff split: split_if_asm)
-            with Suc show ?case by simp
-          qed simp }
-        from this[of "l - k"] `k \<le> l` have "w l i = w k i" by simp }
-      note w_mono = this
-
-      def w' \<equiv> "\<lambda>i. if i \<in> (\<Union>k. J k) then w (LEAST k. i \<in> J k) i else if i \<in> I then (SOME x. x \<in> space (M i)) else undefined"
-      { fix i k assume k: "i \<in> J k"
-        have "w k i = w (LEAST k. i \<in> J k) i"
-          by (intro w_mono Least_le k LeastI[of _ k])
-        then have "w' i = w k i"
-          unfolding w'_def using k by auto }
-      note w'_eq = this
-      have w'_simps1: "\<And>i. i \<notin> I \<Longrightarrow> w' i = undefined"
-        using J by (auto simp: w'_def)
-      have w'_simps2: "\<And>i. i \<notin> (\<Union>k. J k) \<Longrightarrow> i \<in> I \<Longrightarrow> w' i \<in> space (M i)"
-        using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]])
-      { fix i assume "i \<in> I" then have "w' i \<in> space (M i)"
-          using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ }
-      note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this
-
-      have w': "w' \<in> space (Pi\<^sub>M I M)"
-        using w(1) by (auto simp add: Pi_iff extensional_def space_PiM)
-
-      { fix n
-        have "restrict w' (J n) = w n" using w(1)[of n]
-          by (auto simp add: fun_eq_iff space_PiM)
-        with w_ext[of n] obtain x where "x \<in> A n" "restrict x (J n) = restrict w' (J n)"
-          by auto
-        then have "w' \<in> A n" unfolding A_eq using w' by (auto simp: prod_emb_def space_PiM) }
-      then have "w' \<in> (\<Inter>i. A i)" by auto
-      with `(\<Inter>i. A i) = {}` show False by auto
-    qed
-    ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0"
-      using LIMSEQ_INF[of "\<lambda>i. \<mu>G (A i)"] by simp
-  qed fact+
-  then guess \<mu> .. note \<mu> = this
-  show ?thesis
-  proof (subst emeasure_extend_measure_Pair[OF PiM_def, of I M \<mu> J X])
-    from assms show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))"
-      by (simp add: Pi_iff)
-  next
-    fix J X assume J: "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))"
-    then show "emb I J (Pi\<^sub>E J X) \<in> Pow (\<Pi>\<^sub>E i\<in>I. space (M i))"
-      by (auto simp: Pi_iff prod_emb_def dest: sets.sets_into_space)
-    have "emb I J (Pi\<^sub>E J X) \<in> generator"
-      using J `I \<noteq> {}` by (intro generatorI') (auto simp: Pi_iff)
-    then have "\<mu> (emb I J (Pi\<^sub>E J X)) = \<mu>G (emb I J (Pi\<^sub>E J X))"
-      using \<mu> by simp
-    also have "\<dots> = (\<Prod> j\<in>J. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
-      using J  `I \<noteq> {}` by (subst mu_G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff)
-    also have "\<dots> = (\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}.
-      if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
-      using J `I \<noteq> {}` by (intro setprod.mono_neutral_right) (auto simp: M.emeasure_space_1)
-    finally show "\<mu> (emb I J (Pi\<^sub>E J X)) = \<dots>" .
-  next
-    let ?F = "\<lambda>j. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j))"
-    have "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) = (\<Prod>j\<in>J. ?F j)"
-      using X `I \<noteq> {}` by (intro setprod.mono_neutral_right) (auto simp: M.emeasure_space_1)
-    then show "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) =
-      emeasure (Pi\<^sub>M J M) (Pi\<^sub>E J X)"
-      using X by (auto simp add: emeasure_PiM) 
-  next
-    show "positive (sets (Pi\<^sub>M I M)) \<mu>" "countably_additive (sets (Pi\<^sub>M I M)) \<mu>"
-      using \<mu> unfolding sets_PiM_generator by (auto simp: measure_space_def)
-  qed
-qed
+lemma (in product_prob_space) emeasure_PiM_emb:
+  "J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> (\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)) \<Longrightarrow> 
+    emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))"
+  by (subst emeasure_PiM_emb') (auto intro!: emeasure_PiM)
 
 sublocale product_prob_space \<subseteq> P: prob_space "Pi\<^sub>M I M"
 proof
+  have *: "emb I {} {\<lambda>x. undefined} = space (PiM I M)"
+    by (auto simp: prod_emb_def space_PiM)
   show "emeasure (Pi\<^sub>M I M) (space (Pi\<^sub>M I M)) = 1"
-  proof cases
-    assume "I = {}" then show ?thesis by (simp add: space_PiM_empty)
-  next
-    assume "I \<noteq> {}"
-    then obtain i where i: "i \<in> I" by auto
-    then have "emb I {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i)) = (space (Pi\<^sub>M I M))"
-      by (auto simp: prod_emb_def space_PiM)
-    with i show ?thesis
-      using emeasure_PiM_emb_not_empty[of "{i}" "\<lambda>i. space (M i)"]
-      by (simp add: emeasure_PiM emeasure_space_1)
-  qed
-qed
-
-lemma (in product_prob_space) emeasure_PiM_emb:
-  assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
-  shows "emeasure (Pi\<^sub>M I M) (emb I J (Pi\<^sub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))"
-proof cases
-  assume "J = {}"
-  moreover have "emb I {} {\<lambda>x. undefined} = space (Pi\<^sub>M I M)"
-    by (auto simp: space_PiM prod_emb_def)
-  ultimately show ?thesis
-    by (simp add: space_PiM_empty P.emeasure_space_1)
-next
-  assume "J \<noteq> {}" with X show ?thesis
-    by (subst emeasure_PiM_emb_not_empty) (auto simp: emeasure_PiM)
+    using emeasure_PiM_emb[of "{}" "\<lambda>_. {}"] by (simp add: *)
 qed
 
 lemma (in product_prob_space) emeasure_PiM_Collect:
@@ -388,7 +102,6 @@
 qed simp
 
 lemma (in product_prob_space) PiM_eq:
-  assumes "I \<noteq> {}"
   assumes "sets M' = sets (PiM I M)"
   assumes eq: "\<And>J F. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> (\<And>j. j \<in> J \<Longrightarrow> F j \<in> sets (M j)) \<Longrightarrow>
     emeasure M' (prod_emb I M J (\<Pi>\<^sub>E j\<in>J. F j)) = (\<Prod>j\<in>J. emeasure (M j) (F j))"
@@ -400,15 +113,15 @@
     unfolding `sets M' = sets (PiM I M)` by simp
 
   def i \<equiv> "SOME i. i \<in> I"
-  with `I \<noteq> {}` have i: "i \<in> I"
-    by (auto intro: someI_ex)
+  have i: "I \<noteq> {} \<Longrightarrow> i \<in> I"
+    unfolding i_def by (rule someI_ex) auto
 
-  def A \<equiv> "\<lambda>n::nat. prod_emb I M {i} (\<Pi>\<^sub>E j\<in>{i}. space (M i))"
+  def A \<equiv> "\<lambda>n::nat. if I = {} then prod_emb I M {} (\<Pi>\<^sub>E i\<in>{}. {}) else prod_emb I M {i} (\<Pi>\<^sub>E i\<in>{i}. space (M i))"
   then show "range A \<subseteq> prod_algebra I M"
-    by (auto intro!: prod_algebraI i)
+    using prod_algebraI[of "{}" I "\<lambda>i. space (M i)" M] by (auto intro!: prod_algebraI i)
 
   have A_eq: "\<And>i. A i = space (PiM I M)"
-    by (auto simp: prod_emb_def space_PiM Pi_iff A_def i)
+    by (auto simp: prod_emb_def space_PiM PiE_iff A_def i ex_in_conv[symmetric] exI)
   show "(\<Union>i. A i) = (\<Pi>\<^sub>E i\<in>I. space (M i))"
     unfolding A_eq by (auto simp: space_PiM)
   show "\<And>i. emeasure (PiM I M) (A i) \<noteq> \<infinity>"
--- a/src/HOL/Probability/Measure_Space.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Measure_Space.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -555,6 +555,53 @@
     unfolding ereal_minus_eq_minus_iff using finite A0 by auto
 qed
 
+lemma emeasure_INT_decseq_subset:
+  fixes F :: "nat \<Rightarrow> 'a set"
+  assumes I: "I \<noteq> {}" and F: "\<And>i j. i \<in> I \<Longrightarrow> j \<in> I \<Longrightarrow> i \<le> j \<Longrightarrow> F j \<subseteq> F i"
+  assumes F_sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets M"
+    and fin: "\<And>i. i \<in> I \<Longrightarrow> emeasure M (F i) \<noteq> \<infinity>"
+  shows "emeasure M (\<Inter>i\<in>I. F i) = (INF i:I. emeasure M (F i))"
+proof cases
+  assume "finite I"
+  have "(\<Inter>i\<in>I. F i) = F (Max I)"
+    using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F) auto
+  moreover have "(INF i:I. emeasure M (F i)) = emeasure M (F (Max I))"
+    using I \<open>finite I\<close> by (intro antisym INF_lower INF_greatest F emeasure_mono) auto
+  ultimately show ?thesis
+    by simp
+next
+  assume "infinite I"
+  def L \<equiv> "\<lambda>n. LEAST i. i \<in> I \<and> i \<ge> n"
+  have L: "L n \<in> I \<and> n \<le> L n" for n
+    unfolding L_def
+  proof (rule LeastI_ex)
+    show "\<exists>x. x \<in> I \<and> n \<le> x"
+      using \<open>infinite I\<close> finite_subset[of I "{..< n}"]
+      by (rule_tac ccontr) (auto simp: not_le)
+  qed
+  have L_eq[simp]: "i \<in> I \<Longrightarrow> L i = i" for i
+    unfolding L_def by (intro Least_equality) auto
+  have L_mono: "i \<le> j \<Longrightarrow> L i \<le> L j" for i j
+    using L[of j] unfolding L_def by (intro Least_le) (auto simp: L_def)
+
+  have "emeasure M (\<Inter>i. F (L i)) = (INF i. emeasure M (F (L i)))"
+  proof (intro INF_emeasure_decseq[symmetric])
+    show "decseq (\<lambda>i. F (L i))"
+      using L by (intro antimonoI F L_mono) auto
+  qed (insert L fin, auto)
+  also have "\<dots> = (INF i:I. emeasure M (F i))"
+  proof (intro antisym INF_greatest)
+    show "i \<in> I \<Longrightarrow> (INF i. emeasure M (F (L i))) \<le> emeasure M (F i)" for i
+      by (intro INF_lower2[of i]) auto
+  qed (insert L, auto intro: INF_lower)
+  also have "(\<Inter>i. F (L i)) = (\<Inter>i\<in>I. F i)"
+  proof (intro antisym INF_greatest)
+    show "i \<in> I \<Longrightarrow> (\<Inter>i. F (L i)) \<subseteq> F i" for i
+      by (intro INF_lower2[of i]) auto
+  qed (insert L, auto)
+  finally show ?thesis .
+qed
+
 lemma Lim_emeasure_decseq:
   assumes A: "range A \<subseteq> sets M" "decseq A" and fin: "\<And>i. emeasure M (A i) \<noteq> \<infinity>"
   shows "(\<lambda>i. emeasure M (A i)) ----> emeasure M (\<Inter>i. A i)"
--- a/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Nonnegative_Lebesgue_Integration.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -1511,8 +1511,10 @@
 qed
 
 lemma nn_integral_monotone_convergence_INF:
-  assumes f: "decseq f" and [measurable]: "\<And>i. f i \<in> borel_measurable M"
-  assumes fin: "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
+  fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal"
+  assumes f: "\<And>i j x. i \<le> j \<Longrightarrow> x \<in> space M \<Longrightarrow> f j x \<le> f i x"
+    and [measurable]: "\<And>i. f i \<in> borel_measurable M"
+    and fin: "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
   shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
 proof -
   { fix f :: "nat \<Rightarrow> ereal" and j assume "decseq f"
@@ -1523,28 +1525,33 @@
       done }
   note INF_shift = this
 
-  have dec: "\<And>f::nat \<Rightarrow> 'a \<Rightarrow> ereal. decseq f \<Longrightarrow> decseq (\<lambda>j x. max 0 (f (j + i) x))"
-    by (intro antimonoI le_funI max.mono) (auto simp: decseq_def le_fun_def)
+  have dec: "decseq (\<lambda>j x. max 0 (restrict (f (j + i)) (space M) x))"
+    using f by (intro antimonoI le_funI max.mono) (auto simp: decseq_def le_fun_def)
 
-  have "(\<integral>\<^sup>+ x. max 0 (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (INF i. max 0 (f i x)) \<partial>M)"
+  have "(\<integral>\<^sup>+ x. max 0 (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (INF i. max 0 (restrict (f i) (space M) x)) \<partial>M)"
     by (intro nn_integral_cong)
        (simp add: sup_ereal_def[symmetric] sup_INF del: sup_ereal_def)
-  also have "\<dots> = (\<integral>\<^sup>+ x. (INF j. max 0 (f (j + i) x)) \<partial>M)"
+  also have "\<dots> = (\<integral>\<^sup>+ x. (INF j. max 0 (restrict (f (j + i)) (space M) x)) \<partial>M)"
     using f by (intro nn_integral_cong INF_shift antimonoI le_funI max.mono) 
                (auto simp: decseq_def le_fun_def)
-  also have "\<dots> = (INF j. (\<integral>\<^sup>+ x. max 0 (f (j + i) x) \<partial>M))"
+  also have "\<dots> = (INF j. (\<integral>\<^sup>+ x. max 0 (restrict (f (j + i)) (space M) x) \<partial>M))"
   proof (rule nn_integral_monotone_convergence_INF')
-    show "\<And>j. (\<lambda>x. max 0 (f (j + i) x)) \<in> borel_measurable M"
-      by measurable
-    show "(\<integral>\<^sup>+ x. max 0 (f (0 + i) x) \<partial>M) < \<infinity>"
-      using fin by (simp add: nn_integral_max_0)
+    show "(\<lambda>x. max 0 (restrict (f (j + i)) (space M) x)) \<in> borel_measurable M" for j
+      by (subst measurable_cong[where g="\<lambda>x. max 0 (f (j + i) x)"]) measurable
+    show "(\<integral>\<^sup>+ x. max 0 (restrict (f (0 + i)) (space M) x) \<partial>M) < \<infinity>"
+      using fin by (simp add: nn_integral_max_0 cong: nn_integral_cong)
   qed (intro max.cobounded1 dec f)+
-  also have "\<dots> = (INF j. (\<integral>\<^sup>+ x. max 0 (f j x) \<partial>M))"
+  also have "\<dots> = (INF j. (\<integral>\<^sup>+ x. max 0 (restrict (f j) (space M) x) \<partial>M))"
     using f by (intro INF_shift[symmetric] nn_integral_mono antimonoI le_funI max.mono) 
                (auto simp: decseq_def le_fun_def)
-  finally show ?thesis unfolding nn_integral_max_0 .
+  finally show ?thesis unfolding nn_integral_max_0 by (simp cong: nn_integral_cong)
 qed
 
+lemma nn_integral_monotone_convergence_INF_decseq:
+  assumes f: "decseq f" and *: "\<And>i. f i \<in> borel_measurable M" "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>"
+  shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))"
+  using nn_integral_monotone_convergence_INF[of M f i, OF _ *] f by (auto simp: antimono_def le_fun_def)
+
 lemma sup_continuous_nn_integral[order_continuous_intros]:
   assumes f: "\<And>y. sup_continuous (f y)"
   assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M"
@@ -1771,7 +1778,7 @@
   fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ereal" assume "decseq C" "\<And>i. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)"
   then show "(\<lambda>s. \<integral>\<^sup>+x. (INF i. C i) x \<partial>M s) = (INF i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))"
     unfolding INF_apply[abs_def]
-    by (subst nn_integral_monotone_convergence_INF)
+    by (subst nn_integral_monotone_convergence_INF_decseq)
        (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets)
 next
   show "\<And>x. g x \<le> (\<lambda>s. integral\<^sup>N (M s) (f top))"
@@ -1782,7 +1789,7 @@
   fix C assume "\<And>i::nat. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)" "decseq C"
   with bound show "INFIMUM UNIV C \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (INFIMUM UNIV C) < \<infinity>)"
     unfolding INF_apply[abs_def]
-    by (subst nn_integral_monotone_convergence_INF)
+    by (subst nn_integral_monotone_convergence_INF_decseq)
        (auto cong: measurable_cong_sets intro!: borel_measurable_INF
              simp: INF_less_iff simp del: ereal_infty_less(1))
 next
--- a/src/HOL/Probability/Probability.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Probability.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -12,9 +12,7 @@
   Probability_Mass_Function
   Stream_Space
   Embed_Measure
-  Interval_Integral
-  Set_Integral
-  Giry_Monad
 begin
 
 end
+
--- a/src/HOL/Probability/Probability_Measure.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Probability_Measure.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -43,6 +43,16 @@
     by (auto simp: emeasure_distr emeasure_space_1)
 qed
 
+lemma prob_space_distrD:
+  assumes f: "f \<in> measurable M N" and M: "prob_space (distr M N f)" shows "prob_space M"
+proof
+  interpret M!: prob_space "distr M N f" by fact
+  have "f -` space N \<inter> space M = space M"
+    using f[THEN measurable_space] by auto
+  then show "emeasure M (space M) = 1"
+    using M.emeasure_space_1 by (simp add: emeasure_distr[OF f])
+qed
+
 lemma (in prob_space) prob_space: "prob (space M) = 1"
   using emeasure_space_1 unfolding measure_def by (simp add: one_ereal_def)
 
@@ -1152,4 +1162,53 @@
   "finite S \<Longrightarrow> (\<And>s. s \<in> S \<Longrightarrow> 0 \<le> p s) \<Longrightarrow> (\<Sum>s\<in>S. p s) = 1 \<Longrightarrow> prob_space (point_measure S p)"
   by (rule prob_spaceI) (simp add: space_point_measure emeasure_point_measure_finite)
 
+lemma (in prob_space) distr_pair_fst: "distr (N \<Otimes>\<^sub>M M) N fst = N"
+proof (intro measure_eqI)
+  fix A assume A: "A \<in> sets (distr (N \<Otimes>\<^sub>M M) N fst)"
+  from A have "emeasure (distr (N \<Otimes>\<^sub>M M) N fst) A = emeasure (N \<Otimes>\<^sub>M M) (A \<times> space M)"
+    by (auto simp add: emeasure_distr space_pair_measure dest: sets.sets_into_space intro!: arg_cong2[where f=emeasure])
+  with A show "emeasure (distr (N \<Otimes>\<^sub>M M) N fst) A = emeasure N A"
+    by (simp add: emeasure_pair_measure_Times emeasure_space_1)
+qed simp
+
+lemma (in product_prob_space) distr_reorder:
+  assumes "inj_on t J" "t \<in> J \<rightarrow> K" "finite K"
+  shows "distr (PiM K M) (Pi\<^sub>M J (\<lambda>x. M (t x))) (\<lambda>\<omega>. \<lambda>n\<in>J. \<omega> (t n)) = PiM J (\<lambda>x. M (t x))"
+proof (rule product_sigma_finite.PiM_eqI)
+  show "product_sigma_finite (\<lambda>x. M (t x))" ..
+  have "t`J \<subseteq> K" using assms by auto
+  then show [simp]: "finite J"
+    by (rule finite_imageD[OF finite_subset]) fact+
+  fix A assume A: "\<And>i. i \<in> J \<Longrightarrow> A i \<in> sets (M (t i))"
+  moreover have "((\<lambda>\<omega>. \<lambda>n\<in>J. \<omega> (t n)) -` Pi\<^sub>E J A \<inter> space (Pi\<^sub>M K M)) =
+    (\<Pi>\<^sub>E i\<in>K. if i \<in> t`J then A (the_inv_into J t i) else space (M i))"
+    using A A[THEN sets.sets_into_space] \<open>t \<in> J \<rightarrow> K\<close> \<open>inj_on t J\<close>
+    by (subst prod_emb_Pi[symmetric]) (auto simp: space_PiM PiE_iff the_inv_into_f_f prod_emb_def)
+  ultimately show "distr (Pi\<^sub>M K M) (Pi\<^sub>M J (\<lambda>x. M (t x))) (\<lambda>\<omega>. \<lambda>n\<in>J. \<omega> (t n)) (Pi\<^sub>E J A) = (\<Prod>i\<in>J. M (t i) (A i))"
+    using assms
+    apply (subst emeasure_distr)
+    apply (auto intro!: sets_PiM_I_finite simp: Pi_iff)
+    apply (subst emeasure_PiM)
+    apply (auto simp: the_inv_into_f_f \<open>inj_on t J\<close> setprod.reindex[OF \<open>inj_on t J\<close>]
+      if_distrib[where f="emeasure (M _)"] setprod.If_cases emeasure_space_1 Int_absorb1 \<open>t`J \<subseteq> K\<close>)
+    done
+qed simp
+
+lemma (in product_prob_space) distr_restrict:
+  "J \<subseteq> K \<Longrightarrow> finite K \<Longrightarrow> (\<Pi>\<^sub>M i\<in>J. M i) = distr (\<Pi>\<^sub>M i\<in>K. M i) (\<Pi>\<^sub>M i\<in>J. M i) (\<lambda>f. restrict f J)"
+  using distr_reorder[of "\<lambda>x. x" J K] by (simp add: Pi_iff subset_eq)
+
+lemma (in product_prob_space) emeasure_prod_emb[simp]:
+  assumes L: "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^sub>M J M)"
+  shows "emeasure (Pi\<^sub>M L M) (prod_emb L M J X) = emeasure (Pi\<^sub>M J M) X"
+  by (subst distr_restrict[OF L])
+     (simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X)
+
+lemma emeasure_distr_restrict:
+  assumes "I \<subseteq> K" and Q[measurable_cong]: "sets Q = sets (PiM K M)" and A[measurable]: "A \<in> sets (PiM I M)"
+  shows "emeasure (distr Q (PiM I M) (\<lambda>\<omega>. restrict \<omega> I)) A = emeasure Q (prod_emb K M I A)"
+  using \<open>I\<subseteq>K\<close> sets_eq_imp_space_eq[OF Q]
+  by (subst emeasure_distr)
+     (auto simp: measurable_cong_sets[OF Q] prod_emb_def space_PiM[symmetric] intro!: measurable_restrict)
+
 end
--- a/src/HOL/Probability/Projective_Family.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Projective_Family.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -6,342 +6,672 @@
 section {*Projective Family*}
 
 theory Projective_Family
-imports Finite_Product_Measure Probability_Measure
+imports Finite_Product_Measure Giry_Monad
 begin
 
-lemma (in product_prob_space) distr_restrict:
-  assumes "J \<noteq> {}" "J \<subseteq> K" "finite K"
-  shows "(\<Pi>\<^sub>M i\<in>J. M i) = distr (\<Pi>\<^sub>M i\<in>K. M i) (\<Pi>\<^sub>M i\<in>J. M i) (\<lambda>f. restrict f J)" (is "?P = ?D")
-proof (rule measure_eqI_generator_eq)
-  have "finite J" using `J \<subseteq> K` `finite K` by (auto simp add: finite_subset)
-  interpret J: finite_product_prob_space M J proof qed fact
-  interpret K: finite_product_prob_space M K proof qed fact
-
-  let ?J = "{Pi\<^sub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
-  let ?F = "\<lambda>i. \<Pi>\<^sub>E k\<in>J. space (M k)"
-  let ?\<Omega> = "(\<Pi>\<^sub>E k\<in>J. space (M k))"
-  show "Int_stable ?J"
-    by (rule Int_stable_PiE)
-  show "range ?F \<subseteq> ?J" "(\<Union>i. ?F i) = ?\<Omega>"
-    using `finite J` by (auto intro!: prod_algebraI_finite)
-  { fix i show "emeasure ?P (?F i) \<noteq> \<infinity>" by simp }
-  show "?J \<subseteq> Pow ?\<Omega>" by (auto simp: Pi_iff dest: sets.sets_into_space)
-  show "sets (\<Pi>\<^sub>M i\<in>J. M i) = sigma_sets ?\<Omega> ?J" "sets ?D = sigma_sets ?\<Omega> ?J"
-    using `finite J` by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff)
-
-  fix X assume "X \<in> ?J"
-  then obtain E where [simp]: "X = Pi\<^sub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto
-  with `finite J` have X: "X \<in> sets (Pi\<^sub>M J M)"
-    by simp
-
-  have "emeasure ?P X = (\<Prod> i\<in>J. emeasure (M i) (E i))"
-    using E by (simp add: J.measure_times)
-  also have "\<dots> = (\<Prod> i\<in>J. emeasure (M i) (if i \<in> J then E i else space (M i)))"
-    by simp
-  also have "\<dots> = (\<Prod> i\<in>K. emeasure (M i) (if i \<in> J then E i else space (M i)))"
-    using `finite K` `J \<subseteq> K`
-    by (intro setprod.mono_neutral_left) (auto simp: M.emeasure_space_1)
-  also have "\<dots> = emeasure (Pi\<^sub>M K M) (\<Pi>\<^sub>E i\<in>K. if i \<in> J then E i else space (M i))"
-    using E by (simp add: K.measure_times)
-  also have "(\<Pi>\<^sub>E i\<in>K. if i \<in> J then E i else space (M i)) = (\<lambda>f. restrict f J) -` Pi\<^sub>E J E \<inter> (\<Pi>\<^sub>E i\<in>K. space (M i))"
-    using `J \<subseteq> K` sets.sets_into_space E by (force simp: Pi_iff PiE_def split: split_if_asm)
-  finally show "emeasure (Pi\<^sub>M J M) X = emeasure ?D X"
-    using X `J \<subseteq> K` apply (subst emeasure_distr)
-    by (auto intro!: measurable_restrict_subset simp: space_PiM)
+lemma vimage_restrict_preseve_mono:
+  assumes J: "J \<subseteq> I"
+  and sets: "A \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^sub>E i\<in>J. S i)" and ne: "(\<Pi>\<^sub>E i\<in>I. S i) \<noteq> {}"
+  and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i) \<subseteq> (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
+  shows "A \<subseteq> B"
+proof  (intro subsetI)
+  fix x assume "x \<in> A"
+  from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto
+  have "J \<inter> (I - J) = {}" by auto
+  show "x \<in> B"
+  proof cases
+    assume x: "x \<in> (\<Pi>\<^sub>E i\<in>J. S i)"
+    have "merge J (I - J) (x,y) \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^sub>E i\<in>I. S i)"
+      using y x `J \<subseteq> I` PiE_cancel_merge[of "J" "I - J" x y S] \<open>x\<in>A\<close>
+      by (auto simp del: PiE_cancel_merge simp add: Un_absorb1)
+    also have "\<dots> \<subseteq> (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^sub>E i\<in>I. S i)" by fact
+    finally show "x \<in> B"
+      using y x `J \<subseteq> I` PiE_cancel_merge[of "J" "I - J" x y S] \<open>x\<in>A\<close> eq
+      by (auto simp del: PiE_cancel_merge simp add: Un_absorb1)
+  qed (insert \<open>x\<in>A\<close> sets, auto)
 qed
 
-lemma (in product_prob_space) emeasure_prod_emb[simp]:
-  assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^sub>M J M)"
-  shows "emeasure (Pi\<^sub>M L M) (prod_emb L M J X) = emeasure (Pi\<^sub>M J M) X"
-  by (subst distr_restrict[OF L])
-     (simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X)
-
-definition
-  limP :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a measure) \<Rightarrow> ('i set \<Rightarrow> ('i \<Rightarrow> 'a) measure) \<Rightarrow> ('i \<Rightarrow> 'a) measure" where
-  "limP I M P = extend_measure (\<Pi>\<^sub>E i\<in>I. space (M i))
-    {(J, X). (J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))}
-    (\<lambda>(J, X). prod_emb I M J (\<Pi>\<^sub>E j\<in>J. X j))
-    (\<lambda>(J, X). emeasure (P J) (Pi\<^sub>E J X))"
-
-abbreviation "lim\<^sub>P \<equiv> limP"
-
-lemma space_limP[simp]: "space (limP I M P) = space (PiM I M)"
-  by (auto simp add: limP_def space_PiM prod_emb_def intro!: space_extend_measure)
-
-lemma sets_limP[simp]: "sets (limP I M P) = sets (PiM I M)"
-  by (auto simp add: limP_def sets_PiM prod_algebra_def prod_emb_def intro!: sets_extend_measure)
-
-lemma measurable_limP1[simp]: "measurable (limP I M P) M' = measurable (\<Pi>\<^sub>M i\<in>I. M i) M'"
-  unfolding measurable_def by auto
-
-lemma measurable_limP2[simp]: "measurable M' (limP I M P) = measurable M' (\<Pi>\<^sub>M i\<in>I. M i)"
-  unfolding measurable_def by auto
-
 locale projective_family =
-  fixes I::"'i set" and P::"'i set \<Rightarrow> ('i \<Rightarrow> 'a) measure" and M::"('i \<Rightarrow> 'a measure)"
-  assumes projective: "\<And>J H X. J \<noteq> {} \<Longrightarrow> J \<subseteq> H \<Longrightarrow> H \<subseteq> I \<Longrightarrow> finite H \<Longrightarrow> X \<in> sets (PiM J M) \<Longrightarrow>
-     (P H) (prod_emb H M J X) = (P J) X"
-  assumes proj_prob_space: "\<And>J. finite J \<Longrightarrow> prob_space (P J)"
-  assumes proj_space: "\<And>J. finite J \<Longrightarrow> space (P J) = space (PiM J M)"
-  assumes proj_sets: "\<And>J. finite J \<Longrightarrow> sets (P J) = sets (PiM J M)"
+  fixes I :: "'i set" and P :: "'i set \<Rightarrow> ('i \<Rightarrow> 'a) measure" and M :: "'i \<Rightarrow> 'a measure"
+  assumes P: "\<And>J H. J \<subseteq> H \<Longrightarrow> finite H \<Longrightarrow> H \<subseteq> I \<Longrightarrow> P J = distr (P H) (PiM J M) (\<lambda>f. restrict f J)"
+  assumes prob_space_P: "\<And>J. finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> prob_space (P J)"
 begin
 
-lemma emeasure_limP:
-  assumes "finite J"
-  assumes "J \<subseteq> I"
-  assumes A: "\<And>i. i\<in>J \<Longrightarrow> A i \<in> sets (M i)"
-  shows "emeasure (limP J M P) (Pi\<^sub>E J A) = emeasure (P J) (Pi\<^sub>E J A)"
-proof -
-  have "Pi\<^sub>E J (restrict A J) \<subseteq> (\<Pi>\<^sub>E i\<in>J. space (M i))"
-    using sets.sets_into_space[OF A] by (auto simp: PiE_iff) blast
-  hence "emeasure (limP J M P) (Pi\<^sub>E J A) =
-    emeasure (limP J M P) (prod_emb J M J (Pi\<^sub>E J A))"
-    using assms(1-3) sets.sets_into_space by (auto simp add: prod_emb_id PiE_def Pi_def)
-  also have "\<dots> = emeasure (P J) (Pi\<^sub>E J A)"
-  proof (rule emeasure_extend_measure_Pair[OF limP_def])
-    show "positive (sets (limP J M P)) (P J)" unfolding positive_def by auto
-    show "countably_additive (sets (limP J M P)) (P J)" unfolding countably_additive_def
-      by (auto simp: suminf_emeasure proj_sets[OF `finite J`])
-    show "(J \<noteq> {} \<or> J = {}) \<and> finite J \<and> J \<subseteq> J \<and> A \<in> (\<Pi> j\<in>J. sets (M j))"
-      using assms by auto
-    fix K and X::"'i \<Rightarrow> 'a set"
-    show "prod_emb J M K (Pi\<^sub>E K X) \<in> Pow (\<Pi>\<^sub>E i\<in>J. space (M i))"
-      by (auto simp: prod_emb_def)
-    assume JX: "(K \<noteq> {} \<or> J = {}) \<and> finite K \<and> K \<subseteq> J \<and> X \<in> (\<Pi> j\<in>K. sets (M j))"
-    thus "emeasure (P J) (prod_emb J M K (Pi\<^sub>E K X)) = emeasure (P K) (Pi\<^sub>E K X)"
-      using assms
-      apply (cases "J = {}")
-      apply (simp add: prod_emb_id)
-      apply (fastforce simp add: intro!: projective sets_PiM_I_finite)
-      done
-  qed
-  finally show ?thesis .
-qed
+lemma sets_P: "finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> sets (P J) = sets (PiM J M)"
+  by (subst P[of J J]) simp_all
+
+lemma space_P: "finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> space (P J) = space (PiM J M)"
+  using sets_P by (rule sets_eq_imp_space_eq)
 
-lemma limP_finite[simp]:
-  assumes "finite J"
-  assumes "J \<subseteq> I"
-  shows "limP J M P = P J" (is "?P = _")
-proof (rule measure_eqI_generator_eq)
-  let ?J = "{Pi\<^sub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
-  let ?\<Omega> = "(\<Pi>\<^sub>E k\<in>J. space (M k))"
-  interpret prob_space "P J" using proj_prob_space `finite J` by simp
-  show "emeasure ?P (\<Pi>\<^sub>E k\<in>J. space (M k)) \<noteq> \<infinity>" using assms `finite J` by (auto simp: emeasure_limP)
-  show "sets (limP J M P) = sigma_sets ?\<Omega> ?J" "sets (P J) = sigma_sets ?\<Omega> ?J"
-    using `finite J` proj_sets by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff)
-  fix X assume "X \<in> ?J"
-  then obtain E where X: "X = Pi\<^sub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto
-  with `finite J` have "X \<in> sets (limP J M P)" by simp
-  have emb_self: "prod_emb J M J (Pi\<^sub>E J E) = Pi\<^sub>E J E"
-    using E sets.sets_into_space
-    by (auto intro!: prod_emb_PiE_same_index)
-  show "emeasure (limP J M P) X = emeasure (P J) X"
-    unfolding X using E
-    by (intro emeasure_limP assms) simp
-qed (auto simp: Pi_iff dest: sets.sets_into_space intro: Int_stable_PiE)
+lemma not_empty_M: "i \<in> I \<Longrightarrow> space (M i) \<noteq> {}"
+  using prob_space_P[THEN prob_space.not_empty] by (auto simp: space_PiM space_P)
 
-lemma emeasure_fun_emb[simp]:
-  assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" "L \<subseteq> I" and X: "X \<in> sets (PiM J M)"
-  shows "emeasure (limP L M P) (prod_emb L M J X) = emeasure (limP J M P) X"
-  using assms
-  by (subst limP_finite) (auto simp: limP_finite finite_subset projective)
+lemma not_empty: "space (PiM I M) \<noteq> {}"
+  by (simp add: not_empty_M)
 
 abbreviation
-  "emb L K X \<equiv> prod_emb L M K X"
+  "emb L K \<equiv> prod_emb L M K"
 
-lemma prod_emb_injective:
-  assumes "J \<subseteq> L" and sets: "X \<in> sets (Pi\<^sub>M J M)" "Y \<in> sets (Pi\<^sub>M J M)"
-  assumes "emb L J X = emb L J Y"
-  shows "X = Y"
-proof (rule injective_vimage_restrict)
+lemma emb_preserve_mono:
+  assumes "J \<subseteq> L" "L \<subseteq> I" and sets: "X \<in> sets (Pi\<^sub>M J M)" "Y \<in> sets (Pi\<^sub>M J M)"
+  assumes "emb L J X \<subseteq> emb L J Y"
+  shows "X \<subseteq> Y"
+proof (rule vimage_restrict_preseve_mono)
   show "X \<subseteq> (\<Pi>\<^sub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^sub>E i\<in>J. space (M i))"
     using sets[THEN sets.sets_into_space] by (auto simp: space_PiM)
-  have "\<forall>i\<in>L. \<exists>x. x \<in> space (M i)"
-  proof
-    fix i assume "i \<in> L"
-    interpret prob_space "P {i}" using proj_prob_space by simp
-    from not_empty show "\<exists>x. x \<in> space (M i)" by (auto simp add: proj_space space_PiM)
-  qed
-  from bchoice[OF this]
-  show "(\<Pi>\<^sub>E i\<in>L. space (M i)) \<noteq> {}" by (auto simp: PiE_def)
-  show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^sub>E i\<in>L. space (M i)) = (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^sub>E i\<in>L. space (M i))"
-    using `prod_emb L M J X = prod_emb L M J Y` by (simp add: prod_emb_def)
+  show "(\<Pi>\<^sub>E i\<in>L. space (M i)) \<noteq> {}"
+    using \<open>L \<subseteq> I\<close> by (auto simp add: not_empty_M space_PiM[symmetric])
+  show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^sub>E i\<in>L. space (M i)) \<subseteq> (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^sub>E i\<in>L. space (M i))"
+    using `prod_emb L M J X \<subseteq> prod_emb L M J Y` by (simp add: prod_emb_def)
 qed fact
 
-definition generator :: "('i \<Rightarrow> 'a) set set" where
-  "generator = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^sub>M J M))"
+lemma emb_injective:
+  assumes L: "J \<subseteq> L" "L \<subseteq> I" and X: "X \<in> sets (Pi\<^sub>M J M)" and Y: "Y \<in> sets (Pi\<^sub>M J M)"
+  shows "emb L J X = emb L J Y \<Longrightarrow> X = Y"
+  by (intro antisym emb_preserve_mono[OF L X Y] emb_preserve_mono[OF L Y X]) auto
+
+lemma emeasure_P: "J \<subseteq> K \<Longrightarrow> finite K \<Longrightarrow> K \<subseteq> I \<Longrightarrow> X \<in> sets (PiM J M) \<Longrightarrow> P K (emb K J X) = P J X"
+  by (auto intro!: emeasure_distr_restrict[symmetric] simp: P sets_P)
 
-lemma generatorI':
-  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> emb I J X \<in> generator"
-  unfolding generator_def by auto
+inductive_set generator :: "('i \<Rightarrow> 'a) set set" where
+  "finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> emb I J X \<in> generator"
+
+lemma algebra_generator: "algebra (space (PiM I M)) generator"
+  unfolding algebra_iff_Int
+proof (safe elim!: generator.cases)
+  fix J X assume J: "finite J" "J \<subseteq> I" and X: "X \<in> sets (PiM J M)"
+
+  from X[THEN sets.sets_into_space] J show "x \<in> emb I J X \<Longrightarrow> x \<in> space (PiM I M)" for x
+    by (auto simp: prod_emb_def space_PiM)
 
-lemma algebra_generator:
-  assumes "I \<noteq> {}" shows "algebra (\<Pi>\<^sub>E i\<in>I. space (M i)) generator" (is "algebra ?\<Omega> ?G")
-  unfolding algebra_def algebra_axioms_def ring_of_sets_iff
-proof (intro conjI ballI)
-  let ?G = generator
-  show "?G \<subseteq> Pow ?\<Omega>"
-    by (auto simp: generator_def prod_emb_def)
-  from `I \<noteq> {}` obtain i where "i \<in> I" by auto
-  then show "{} \<in> ?G"
-    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\<lambda>i. {}"]
-             simp: sigma_sets.Empty generator_def prod_emb_def)
-  from `i \<in> I` show "?\<Omega> \<in> ?G"
-    by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^sub>E {i} (\<lambda>i. space (M i))"]
-             simp: generator_def prod_emb_def)
-  fix A assume "A \<in> ?G"
-  then obtain JA XA where XA: "JA \<noteq> {}" "finite JA" "JA \<subseteq> I" "XA \<in> sets (Pi\<^sub>M JA M)" and A: "A = emb I JA XA"
-    by (auto simp: generator_def)
-  fix B assume "B \<in> ?G"
-  then obtain JB XB where XB: "JB \<noteq> {}" "finite JB" "JB \<subseteq> I" "XB \<in> sets (Pi\<^sub>M JB M)" and B: "B = emb I JB XB"
-    by (auto simp: generator_def)
-  let ?RA = "emb (JA \<union> JB) JA XA"
-  let ?RB = "emb (JA \<union> JB) JB XB"
-  have *: "A - B = emb I (JA \<union> JB) (?RA - ?RB)" "A \<union> B = emb I (JA \<union> JB) (?RA \<union> ?RB)"
-    using XA A XB B by auto
-  show "A - B \<in> ?G" "A \<union> B \<in> ?G"
-    unfolding * using XA XB by (safe intro!: generatorI') auto
-qed
+  have "emb I J (space (PiM J M) - X) \<in> generator"
+    by (intro generator.intros J sets.Diff sets.top X)
+  with J show "space (Pi\<^sub>M I M) - emb I J X \<in> generator"
+    by (simp add: space_PiM prod_emb_PiE)
+  
+  fix K Y assume K: "finite K" "K \<subseteq> I" and Y: "Y \<in> sets (PiM K M)"
+  have "emb I (J \<union> K) (emb (J \<union> K) J X) \<inter> emb I (J \<union> K) (emb (J \<union> K) K Y) \<in> generator"
+    unfolding prod_emb_Int[symmetric]
+    by (intro generator.intros sets.Int measurable_prod_emb) (auto intro!: J K X Y)
+  with J K X Y show "emb I J X \<inter> emb I K Y \<in> generator"
+    by simp
+qed (force simp: generator.simps prod_emb_empty[symmetric])
 
-lemma sets_PiM_generator:
-  "sets (PiM I M) = sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) generator"
-proof cases
-  assume "I = {}" then show ?thesis
-    unfolding generator_def
-    by (auto simp: sets_PiM_empty sigma_sets_empty_eq cong: conj_cong)
-next
-  assume "I \<noteq> {}"
-  show ?thesis
-  proof
-    show "sets (Pi\<^sub>M I M) \<subseteq> sigma_sets (\<Pi>\<^sub>E i\<in>I. space (M i)) generator"
-      unfolding sets_PiM
-    proof (safe intro!: sigma_sets_subseteq)
-      fix A assume "A \<in> prod_algebra I M" with `I \<noteq> {}` show "A \<in> generator"
-        by (auto intro!: generatorI' sets_PiM_I_finite elim!: prod_algebraE)
-    qed
-  qed (auto simp: generator_def space_PiM[symmetric] intro!: sets.sigma_sets_subset)
-qed
+interpretation generator!: algebra "space (PiM I M)" generator
+  by (rule algebra_generator)
 
-lemma generatorI:
-  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> generator"
-  unfolding generator_def by auto
+lemma sets_PiM_generator: "sets (PiM I M) = sigma_sets (space (PiM I M)) generator"
+proof (intro antisym sets.sigma_sets_subset)
+  show "sets (PiM I M) \<subseteq> sigma_sets (space (PiM I M)) generator"
+    unfolding sets_PiM_single space_PiM[symmetric]
+  proof (intro sigma_sets_mono', safe)
+    fix i A assume "i \<in> I" and A: "A \<in> sets (M i)"
+    then have "{f \<in> space (Pi\<^sub>M I M). f i \<in> A} = emb I {i} (\<Pi>\<^sub>E j\<in>{i}. A)"
+      by (auto simp: prod_emb_def space_PiM restrict_def Pi_iff PiE_iff extensional_def)
+    with \<open>i\<in>I\<close> A show "{f \<in> space (Pi\<^sub>M I M). f i \<in> A} \<in> generator"
+      by (auto intro!: generator.intros sets_PiM_I_finite)
+  qed
+qed (auto elim!: generator.cases)
 
 definition mu_G ("\<mu>G") where
-  "\<mu>G A =
-    (THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^sub>M J M). A = emb I J X \<longrightarrow> x = emeasure (limP J M P) X))"
+  "\<mu>G A = (THE x. \<forall>J\<subseteq>I. finite J \<longrightarrow> (\<forall>X\<in>sets (Pi\<^sub>M J M). A = emb I J X \<longrightarrow> x = emeasure (P J) X))"
+
+definition lim :: "('i \<Rightarrow> 'a) measure" where
+  "lim = extend_measure (space (PiM I M)) generator (\<lambda>x. x) \<mu>G"
+
+lemma space_lim[simp]: "space lim = space (PiM I M)"
+  using generator.space_closed
+  unfolding lim_def by (intro space_extend_measure) simp
+
+lemma sets_lim[simp, measurable]: "sets lim = sets (PiM I M)"
+  using generator.space_closed by (simp add: lim_def sets_PiM_generator sets_extend_measure)
 
 lemma mu_G_spec:
-  assumes J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^sub>M J M)"
-  shows "\<mu>G A = emeasure (limP J M P) X"
+  assumes J: "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^sub>M J M)"
+  shows "\<mu>G (emb I J X) = emeasure (P J) X"
   unfolding mu_G_def
 proof (intro the_equality allI impI ballI)
-  fix K Y assume K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "A = emb I K Y" "Y \<in> sets (Pi\<^sub>M K M)"
-  have "emeasure (limP K M P) Y = emeasure (limP (K \<union> J) M P) (emb (K \<union> J) K Y)"
-    using K J by (simp del: limP_finite)
+  fix K Y assume K: "finite K" "K \<subseteq> I" "Y \<in> sets (Pi\<^sub>M K M)"
+    and [simp]: "emb I J X = emb I K Y"
+  have "emeasure (P K) Y = emeasure (P (K \<union> J)) (emb (K \<union> J) K Y)"
+    using K J by (simp add: emeasure_P)
   also have "emb (K \<union> J) K Y = emb (K \<union> J) J X"
-    using K J by (simp add: prod_emb_injective[of "K \<union> J" I])
-  also have "emeasure (limP (K \<union> J) M P) (emb (K \<union> J) J X) = emeasure (limP J M P) X"
-    using K J by (simp del: limP_finite)
-  finally show "emeasure (limP J M P) X = emeasure (limP K M P) Y" ..
+    using K J by (simp add: emb_injective[of "K \<union> J" I])
+  also have "emeasure (P (K \<union> J)) (emb (K \<union> J) J X) = emeasure (P J) X"
+    using K J by (subst emeasure_P) simp_all
+  finally show "emeasure (P J) X = emeasure (P K) Y" ..
 qed (insert J, force)
 
-lemma mu_G_eq:
-  "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = emeasure (limP J M P) X"
-  by (intro mu_G_spec) auto
+lemma positive_mu_G: "positive generator \<mu>G"
+proof -
+  show ?thesis
+  proof (safe intro!: positive_def[THEN iffD2])
+    obtain J where "finite J" "J \<subseteq> I" by auto
+    then have "\<mu>G (emb I J {}) = 0"
+      by (subst mu_G_spec) auto
+    then show "\<mu>G {} = 0" by simp
+  qed (auto simp: mu_G_spec elim!: generator.cases)
+qed
 
-lemma generator_Ex:
-  assumes *: "A \<in> generator"
-  shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^sub>M J M) \<and> A = emb I J X \<and> \<mu>G A = emeasure (limP J M P) X"
-proof -
-  from * obtain J X where J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^sub>M J M)"
-    unfolding generator_def by auto
-  with mu_G_spec[OF this] show ?thesis by (auto simp del: limP_finite)
+lemma additive_mu_G: "additive generator \<mu>G"
+proof (safe intro!: additive_def[THEN iffD2] elim!: generator.cases)
+  fix J X K Y assume J: "finite J" "J \<subseteq> I" "X \<in> sets (PiM J M)"
+    and K: "finite K" "K \<subseteq> I" "Y \<in> sets (PiM K M)"
+    and "emb I J X \<inter> emb I K Y = {}"
+  then have JK_disj: "emb (J \<union> K) J X \<inter> emb (J \<union> K) K Y = {}"
+    by (intro emb_injective[of "J \<union> K" I _ "{}"]) (auto simp: sets.Int prod_emb_Int)
+  have "\<mu>G (emb I J X \<union> emb I K Y) = \<mu>G (emb I (J \<union> K) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y))"
+    using J K by simp
+  also have "\<dots> = emeasure (P (J \<union> K)) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)"
+    using J K by (simp add: mu_G_spec sets.Un del: prod_emb_Un)
+  also have "\<dots> = \<mu>G (emb I J X) + \<mu>G (emb I K Y)"
+    using J K JK_disj by (simp add: plus_emeasure[symmetric] mu_G_spec emeasure_P sets_P)
+  finally show "\<mu>G (emb I J X \<union> emb I K Y) = \<mu>G (emb I J X) + \<mu>G (emb I K Y)" .
 qed
 
-lemma generatorE:
-  assumes A: "A \<in> generator"
-  obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^sub>M J M)" "emb I J X = A" "\<mu>G A = emeasure (limP J M P) X"
-  using generator_Ex[OF A] by atomize_elim auto
-
-lemma merge_sets:
-  "J \<inter> K = {} \<Longrightarrow> A \<in> sets (Pi\<^sub>M (J \<union> K) M) \<Longrightarrow> x \<in> space (Pi\<^sub>M J M) \<Longrightarrow> (\<lambda>y. merge J K (x,y)) -` A \<inter> space (Pi\<^sub>M K M) \<in> sets (Pi\<^sub>M K M)"
-  by simp
-
-lemma merge_emb:
-  assumes "K \<subseteq> I" "J \<subseteq> I" and y: "y \<in> space (Pi\<^sub>M J M)"
-  shows "((\<lambda>x. merge J (I - J) (y, x)) -` emb I K X \<inter> space (Pi\<^sub>M I M)) =
-    emb I (K - J) ((\<lambda>x. merge J (K - J) (y, x)) -` emb (J \<union> K) K X \<inter> space (Pi\<^sub>M (K - J) M))"
+lemma emeasure_lim:
+  assumes JX: "finite J" "J \<subseteq> I" "X \<in> sets (PiM J M)"
+  assumes cont: "\<And>J X. (\<And>i. J i \<subseteq> I) \<Longrightarrow> incseq J \<Longrightarrow> (\<And>i. finite (J i)) \<Longrightarrow> (\<And>i. X i \<in> sets (PiM (J i) M)) \<Longrightarrow>
+    decseq (\<lambda>i. emb I (J i) (X i)) \<Longrightarrow> 0 < (INF i. P (J i) (X i)) \<Longrightarrow> (\<Inter>i. emb I (J i) (X i)) \<noteq> {}"
+  shows "emeasure lim (emb I J X) = P J X"
 proof -
-  have [simp]: "\<And>x J K L. merge J K (y, restrict x L) = merge J (K \<inter> L) (y, x)"
-    by (auto simp: restrict_def merge_def)
-  have [simp]: "\<And>x J K L. restrict (merge J K (y, x)) L = merge (J \<inter> L) (K \<inter> L) (y, x)"
-    by (auto simp: restrict_def merge_def)
-  have [simp]: "(I - J) \<inter> K = K - J" using `K \<subseteq> I` `J \<subseteq> I` by auto
-  have [simp]: "(K - J) \<inter> (K \<union> J) = K - J" by auto
-  have [simp]: "(K - J) \<inter> K = K - J" by auto
-  from y `K \<subseteq> I` `J \<subseteq> I` show ?thesis
-    by (simp split: split_merge add: prod_emb_def Pi_iff PiE_def extensional_merge_sub set_eq_iff space_PiM)
-       auto
-qed
-
-lemma positive_mu_G:
-  assumes "I \<noteq> {}"
-  shows "positive generator \<mu>G"
-proof -
-  interpret G!: algebra "\<Pi>\<^sub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
+  have "\<exists>\<mu>. (\<forall>s\<in>generator. \<mu> s = \<mu>G s) \<and>
+    measure_space (space (PiM I M)) (sigma_sets (space (PiM I M)) generator) \<mu>"
+  proof (rule generator.caratheodory_empty_continuous[OF positive_mu_G additive_mu_G])
+    show "\<And>A. A \<in> generator \<Longrightarrow> \<mu>G A \<noteq> \<infinity>"
+    proof (clarsimp elim!: generator.cases simp: mu_G_spec del: notI)
+      fix J assume "finite J" "J \<subseteq> I"
+      then interpret prob_space "P J" by (rule prob_space_P)
+      show "\<And>X. X \<in> sets (Pi\<^sub>M J M) \<Longrightarrow> emeasure (P J) X \<noteq> \<infinity>"
+        by simp
+    qed
+  next
+    fix A assume "range A \<subseteq> generator" and "decseq A" "(\<Inter>i. A i) = {}"
+    then have "\<forall>i. \<exists>J X. A i = emb I J X \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (PiM J M)"
+      unfolding image_subset_iff generator.simps by blast
+    then obtain J X where A: "\<And>i. A i = emb I (J i) (X i)"
+      and *: "\<And>i. finite (J i)" "\<And>i. J i \<subseteq> I" "\<And>i. X i \<in> sets (PiM (J i) M)"
+      by metis
+    have "(INF i. P (J i) (X i)) = 0"
+    proof (rule ccontr)
+      assume INF_P: "(INF i. P (J i) (X i)) \<noteq> 0"
+      have "(\<Inter>i. emb I (\<Union>n\<le>i. J n) (emb (\<Union>n\<le>i. J n) (J i) (X i))) \<noteq> {}"
+      proof (rule cont)
+        show "decseq (\<lambda>i. emb I (\<Union>n\<le>i. J n) (emb (\<Union>n\<le>i. J n) (J i) (X i)))"
+          using * \<open>decseq A\<close> by (subst prod_emb_trans) (auto simp: A[abs_def])
+        show "0 < (INF i. P (\<Union>n\<le>i. J n) (emb (\<Union>n\<le>i. J n) (J i) (X i)))"
+           using * INF_P by (subst emeasure_P) (auto simp: less_le intro!: INF_greatest)
+        show "incseq (\<lambda>i. \<Union>n\<le>i. J n)"
+          by (force simp: incseq_def)
+      qed (insert *, auto)
+      with \<open>(\<Inter>i. A i) = {}\<close> * show False
+        by (subst (asm) prod_emb_trans) (auto simp: A[abs_def])
+    qed
+    moreover have "(\<lambda>i. P (J i) (X i)) ----> (INF i. P (J i) (X i))"
+    proof (intro LIMSEQ_INF antimonoI)
+      fix x y :: nat assume "x \<le> y"
+      have "P (J y \<union> J x) (emb (J y \<union> J x) (J y) (X y)) \<le> P (J y \<union> J x) (emb (J y \<union> J x) (J x) (X x))"
+        using \<open>decseq A\<close>[THEN decseqD, OF \<open>x\<le>y\<close>] *
+        by (auto simp: A sets_P del: subsetI intro!: emeasure_mono  \<open>x \<le> y\<close>
+              emb_preserve_mono[of "J y \<union> J x" I, where X="emb (J y \<union> J x) (J y) (X y)"])
+      then show "P (J y) (X y) \<le> P (J x) (X x)"
+        using * by (simp add: emeasure_P)
+    qed
+    ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0"
+      by (auto simp: A[abs_def] mu_G_spec *)
+  qed
+  then obtain \<mu> where eq: "\<forall>s\<in>generator. \<mu> s = \<mu>G s"
+    and ms: "measure_space (space (PiM I M)) (sets (PiM I M)) \<mu>"
+    by (metis sets_PiM_generator)
   show ?thesis
-  proof (intro positive_def[THEN iffD2] conjI ballI)
-    from generatorE[OF G.empty_sets] guess J X . note this[simp]
-    have "X = {}"
-      by (rule prod_emb_injective[of J I]) simp_all
-    then show "\<mu>G {} = 0" by simp
-  next
-    fix A assume "A \<in> generator"
-    from generatorE[OF this] guess J X . note this[simp]
-    show "0 \<le> \<mu>G A" by (simp add: emeasure_nonneg)
-  qed
-qed
-
-lemma additive_mu_G:
-  assumes "I \<noteq> {}"
-  shows "additive generator \<mu>G"
-proof -
-  interpret G!: algebra "\<Pi>\<^sub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
-  show ?thesis
-  proof (intro additive_def[THEN iffD2] ballI impI)
-    fix A assume "A \<in> generator" with generatorE guess J X . note J = this
-    fix B assume "B \<in> generator" with generatorE guess K Y . note K = this
-    assume "A \<inter> B = {}"
-    have JK: "J \<union> K \<noteq> {}" "J \<union> K \<subseteq> I" "finite (J \<union> K)"
-      using J K by auto
-    have JK_disj: "emb (J \<union> K) J X \<inter> emb (J \<union> K) K Y = {}"
-      apply (rule prod_emb_injective[of "J \<union> K" I])
-      apply (insert `A \<inter> B = {}` JK J K)
-      apply (simp_all add: sets.Int prod_emb_Int)
-      done
-    have AB: "A = emb I (J \<union> K) (emb (J \<union> K) J X)" "B = emb I (J \<union> K) (emb (J \<union> K) K Y)"
-      using J K by simp_all
-    then have "\<mu>G (A \<union> B) = \<mu>G (emb I (J \<union> K) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y))"
-      by simp
-    also have "\<dots> = emeasure (limP (J \<union> K) M P) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)"
-      using JK J(1, 4) K(1, 4) by (simp add: mu_G_eq sets.Un del: prod_emb_Un)
-    also have "\<dots> = \<mu>G A + \<mu>G B"
-      using J K JK_disj by (simp add: plus_emeasure[symmetric] del: limP_finite)
-    finally show "\<mu>G (A \<union> B) = \<mu>G A + \<mu>G B" .
+  proof (subst emeasure_extend_measure[OF lim_def])
+    show "A \<in> generator \<Longrightarrow> \<mu> A = \<mu>G A" for A
+      using eq by simp
+    show "positive (sets lim) \<mu>" "countably_additive (sets lim) \<mu>"
+      using ms by (auto simp add: measure_space_def)
+    show "(\<lambda>x. x) ` generator \<subseteq> Pow (space (Pi\<^sub>M I M))"
+      using generator.space_closed by simp
+    show "emb I J X \<in> generator" "\<mu>G (emb I J X) = emeasure (P J) X"
+      using JX by (auto intro: generator.intros simp: mu_G_spec)
   qed
 qed
 
 end
 
 sublocale product_prob_space \<subseteq> projective_family I "\<lambda>J. PiM J M" M
-proof (simp add: projective_family_def, safe)
-  fix J::"'i set" assume [simp]: "finite J"
-  interpret f: finite_product_prob_space M J proof qed fact
-  show "prob_space (Pi\<^sub>M J M)"
-  proof
-    show "emeasure (Pi\<^sub>M J M) (space (Pi\<^sub>M J M)) = 1"
-      by (simp add: space_PiM emeasure_PiM emeasure_space_1)
+  unfolding projective_family_def
+proof (intro conjI allI impI distr_restrict)
+  show "\<And>J. finite J \<Longrightarrow> prob_space (Pi\<^sub>M J M)"
+    by (intro prob_spaceI) (simp add: space_PiM emeasure_PiM emeasure_space_1)
+qed auto
+
+
+txt \<open> Proof due to Ionescu Tulcea. \<close>
+
+locale Ionescu_Tulcea =
+  fixes P :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> 'a measure" and M :: "nat \<Rightarrow> 'a measure"
+  assumes P[measurable]: "\<And>i. P i \<in> measurable (PiM {0..<i} M) (subprob_algebra (M i))"
+  assumes prob_space_P: "\<And>i x. x \<in> space (PiM {0..<i} M) \<Longrightarrow> prob_space (P i x)"
+begin
+
+lemma non_empty[simp]: "space (M i) \<noteq> {}"
+proof (induction i rule: less_induct)
+  case (less i)
+  then obtain x where "\<And>j. j < i \<Longrightarrow> x j \<in> space (M j)"
+    unfolding ex_in_conv[symmetric] by metis
+  then have *: "restrict x {0..<i} \<in> space (PiM {0..<i} M)"
+    by (auto simp: space_PiM PiE_iff)
+  then interpret prob_space "P i (restrict x {0..<i})"
+    by (rule prob_space_P)
+  show ?case
+    using not_empty subprob_measurableD(1)[OF P, OF *] by simp
+qed
+
+lemma space_PiM_not_empty[simp]: "space (PiM UNIV M) \<noteq> {}"
+  unfolding space_PiM_empty_iff by auto
+
+lemma space_P: "x \<in> space (PiM {0..<n} M) \<Longrightarrow> space (P n x) = space (M n)"
+  by (simp add: P[THEN subprob_measurableD(1)])
+
+lemma sets_P[measurable_cong]: "x \<in> space (PiM {0..<n} M) \<Longrightarrow> sets (P n x) = sets (M n)"
+  by (simp add: P[THEN subprob_measurableD(2)])
+
+definition eP :: "nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) measure" where
+  "eP n \<omega> = distr (P n \<omega>) (PiM {0..<Suc n} M) (fun_upd \<omega> n)"
+
+lemma measurable_eP[measurable]:
+  "eP n \<in> measurable (PiM {0..< n} M) (subprob_algebra (PiM {0..<Suc n} M))"
+  by (auto simp: eP_def[abs_def] measurable_split_conv
+           intro!: measurable_fun_upd[where J="{0..<n}"] measurable_distr2[OF _ P])
+
+lemma space_eP:
+  "x \<in> space (PiM {0..<n} M) \<Longrightarrow> space (eP n x) = space (PiM {0..<Suc n} M)"
+  by (simp add: measurable_eP[THEN subprob_measurableD(1)])
+
+lemma sets_eP[measurable]:
+  "x \<in> space (PiM {0..<n} M) \<Longrightarrow> sets (eP n x) = sets (PiM {0..<Suc n} M)"
+  by (simp add: measurable_eP[THEN subprob_measurableD(2)])
+
+lemma prob_space_eP: "x \<in> space (PiM {0..<n} M) \<Longrightarrow> prob_space (eP n x)"
+  unfolding eP_def
+  by (intro prob_space.prob_space_distr prob_space_P measurable_fun_upd[where J="{0..<n}"]) auto
+
+lemma nn_integral_eP:
+  "\<omega> \<in> space (PiM {0..<n} M) \<Longrightarrow> f \<in> borel_measurable (PiM {0..<Suc n} M) \<Longrightarrow>
+    (\<integral>\<^sup>+x. f x \<partial>eP n \<omega>) = (\<integral>\<^sup>+x. f (\<omega>(n := x)) \<partial>P n \<omega>)"
+  unfolding eP_def
+  by (subst nn_integral_distr) (auto intro!: measurable_fun_upd[where J="{0..<n}"] simp: space_PiM PiE_iff)
+
+lemma emeasure_eP:
+  assumes \<omega>[simp]: "\<omega> \<in> space (PiM {0..<n} M)" and A[measurable]: "A \<in> sets (PiM {0..<Suc n} M)"
+  shows "eP n \<omega> A = P n \<omega> ((\<lambda>x. \<omega>(n := x)) -` A \<inter> space (M n))"
+  using nn_integral_eP[of \<omega> n "indicator A"]
+  apply (simp add: sets_eP nn_integral_indicator[symmetric] sets_P del: nn_integral_indicator)
+  apply (subst nn_integral_indicator[symmetric])
+  using measurable_sets[OF measurable_fun_upd[OF _ measurable_const[OF \<omega>] measurable_id] A, of n]
+  apply (auto simp add: sets_P atLeastLessThanSuc space_P simp del: nn_integral_indicator
+     intro!: nn_integral_cong split: split_indicator)
+  done
+  
+
+primrec C :: "nat \<Rightarrow> nat \<Rightarrow> (nat \<Rightarrow> 'a) \<Rightarrow> (nat \<Rightarrow> 'a) measure" where
+  "C n 0 \<omega> = return (PiM {0..<n} M) \<omega>"
+| "C n (Suc m) \<omega> = C n m \<omega> \<guillemotright>= eP (n + m)"
+
+lemma measurable_C[measurable]:
+  "C n m \<in> measurable (PiM {0..<n} M) (subprob_algebra (PiM {0..<n + m} M))"
+  by (induction m) auto
+
+lemma space_C:
+  "x \<in> space (PiM {0..<n} M) \<Longrightarrow> space (C n m x) = space (PiM {0..<n + m} M)"
+  by (simp add: measurable_C[THEN subprob_measurableD(1)])
+
+lemma sets_C[measurable_cong]:
+  "x \<in> space (PiM {0..<n} M) \<Longrightarrow> sets (C n m x) = sets (PiM {0..<n + m} M)"
+  by (simp add: measurable_C[THEN subprob_measurableD(2)])
+
+lemma prob_space_C: "x \<in> space (PiM {0..<n} M) \<Longrightarrow> prob_space (C n m x)"
+proof (induction m)
+  case (Suc m) then show ?case
+    by (auto intro!: prob_space.prob_space_bind[where S="PiM {0..<Suc (n + m)} M"]
+             simp: space_C prob_space_eP)
+qed (auto intro!: prob_space_return simp: space_PiM)
+
+lemma split_C:
+  assumes \<omega>: "\<omega> \<in> space (PiM {0..<n} M)" shows "(C n m \<omega> \<guillemotright>= C (n + m) l) = C n (m + l) \<omega>"
+proof (induction l)
+  case 0
+  with \<omega> show ?case
+    by (simp add: bind_return_distr' prob_space_C[THEN prob_space.not_empty]
+                  distr_cong[OF refl sets_C[symmetric, OF \<omega>]])
+next
+  case (Suc l) with \<omega> show ?case
+    by (simp add: bind_assoc[symmetric, OF _ measurable_eP]) (simp add: ac_simps)
+qed
+
+lemma nn_integral_C:
+  assumes "m \<le> m'" and f[measurable]: "f \<in> borel_measurable (PiM {0..<n+m} M)"
+    and nonneg: "\<And>x. x \<in> space (PiM {0..<n+m} M) \<Longrightarrow> 0 \<le> f x"
+    and x: "x \<in> space (PiM {0..<n} M)"
+  shows "(\<integral>\<^sup>+x. f x \<partial>C n m x) = (\<integral>\<^sup>+x. f (restrict x {0..<n+m}) \<partial>C n m' x)"
+  using \<open>m \<le> m'\<close>
+proof (induction rule: dec_induct)
+  case (step i)
+  let ?E = "\<lambda>x. f (restrict x {0..<n + m})" and ?C = "\<lambda>i f. \<integral>\<^sup>+x. f x \<partial>C n i x"
+  from \<open>m\<le>i\<close> x have "?C i ?E = ?C (Suc i) ?E"
+    by (auto simp: nn_integral_bind[where B="PiM {0 ..< Suc (n + i)} M"] space_C nn_integral_eP
+             intro!: nn_integral_cong)
+       (simp add: space_PiM PiE_iff  nonneg prob_space.emeasure_space_1[OF prob_space_P])
+  with step show ?case by (simp del: restrict_apply)
+qed (auto simp: space_PiM space_C[OF x] simp del: restrict_apply intro!: nn_integral_cong)
+
+lemma emeasure_C:
+  assumes "m \<le> m'" and A[measurable]: "A \<in> sets (PiM {0..<n+m} M)" and [simp]: "x \<in> space (PiM {0..<n} M)"
+  shows "emeasure (C n m' x) (prod_emb {0..<n + m'} M {0..<n+m} A) = emeasure (C n m x) A"
+  using assms
+  by (subst (1 2) nn_integral_indicator[symmetric])
+     (auto intro!: nn_integral_cong split: split_indicator simp del: nn_integral_indicator
+           simp: nn_integral_C[of m m' _ n] prod_emb_iff space_PiM PiE_iff sets_C space_C)
+
+lemma distr_C:
+  assumes "m \<le> m'" and [simp]: "x \<in> space (PiM {0..<n} M)"
+  shows "C n m x = distr (C n m' x) (PiM {0..<n+m} M) (\<lambda>x. restrict x {0..<n+m})"
+proof (rule measure_eqI)
+  fix A assume "A \<in> sets (C n m x)"
+  with \<open>m \<le> m'\<close> show "emeasure (C n m x) A = emeasure (distr (C n m' x) (Pi\<^sub>M {0..<n + m} M) (\<lambda>x. restrict x {0..<n + m})) A"
+    by (subst emeasure_C[symmetric, OF \<open>m \<le> m'\<close>]) (auto intro!: emeasure_distr_restrict[symmetric] simp: sets_C)
+qed (simp add: sets_C)
+
+definition up_to :: "nat set \<Rightarrow> nat" where
+  "up_to J = (LEAST n. \<forall>i\<ge>n. i \<notin> J)"
+
+lemma up_to_less: "finite J \<Longrightarrow> i \<in> J \<Longrightarrow> i < up_to J"
+  unfolding up_to_def
+  by (rule LeastI2[of _ "Suc (Max J)"]) (auto simp: Suc_le_eq not_le[symmetric])
+
+lemma up_to_iff: "finite J \<Longrightarrow> up_to J \<le> n \<longleftrightarrow> (\<forall>i\<in>J. i < n)"
+proof safe
+  show "finite J \<Longrightarrow> up_to J \<le> n \<Longrightarrow> i \<in> J \<Longrightarrow> i < n" for i
+    using up_to_less[of J i] by auto
+qed (auto simp: up_to_def intro!: Least_le)
+
+lemma up_to_iff_Ico: "finite J \<Longrightarrow> up_to J \<le> n \<longleftrightarrow> J \<subseteq> {0..<n}"
+  by (auto simp: up_to_iff)
+
+lemma up_to: "finite J \<Longrightarrow> J \<subseteq> {0..< up_to J}"
+  by (auto simp: up_to_less)
+
+lemma up_to_mono: "J \<subseteq> H \<Longrightarrow> finite H \<Longrightarrow> up_to J \<le> up_to H"
+  by (auto simp add: up_to_iff finite_subset up_to_less)
+
+definition CI :: "nat set \<Rightarrow> (nat \<Rightarrow> 'a) measure" where
+  "CI J = distr (C 0 (up_to J) (\<lambda>x. undefined)) (PiM J M) (\<lambda>f. restrict f J)"
+
+sublocale PF!: projective_family UNIV CI
+  unfolding projective_family_def
+proof safe
+  show "finite J \<Longrightarrow> prob_space (CI J)" for J
+    using up_to[of J] unfolding CI_def
+    by (intro prob_space.prob_space_distr prob_space_C measurable_restrict) auto
+  note measurable_cong_sets[OF sets_C, simp]
+  have [simp]: "J \<subseteq> H \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^sub>M H M) (Pi\<^sub>M J M)" for H J
+    by (auto intro!: measurable_restrict)
+
+  show "J \<subseteq> H \<Longrightarrow> finite H \<Longrightarrow> CI J = distr (CI H) (Pi\<^sub>M J M) (\<lambda>f. restrict f J)" for J H
+    by (simp add: CI_def distr_C[OF up_to_mono[of J H]] up_to up_to_mono distr_distr comp_def
+                  inf.absorb2 finite_subset)
+qed
+
+lemma emeasure_CI':
+  "finite J \<Longrightarrow> X \<in> sets (PiM J M) \<Longrightarrow> CI J X = C 0 (up_to J) (\<lambda>_. undefined) (PF.emb {0..<up_to J} J X)"
+  unfolding CI_def using up_to[of J] by (rule emeasure_distr_restrict) (auto simp: sets_C)
+
+lemma emeasure_CI:
+  "J \<subseteq> {0..<n} \<Longrightarrow> X \<in> sets (PiM J M) \<Longrightarrow> CI J X = C 0 n (\<lambda>_. undefined) (PF.emb {0..<n} J X)"
+  apply (subst emeasure_CI', simp_all add: finite_subset)
+  apply (subst emeasure_C[symmetric, of "up_to J" n])
+  apply (auto simp: finite_subset up_to_iff_Ico up_to_less)
+  apply (subst prod_emb_trans)
+  apply (auto simp: up_to_less finite_subset up_to_iff_Ico)
+  done
+
+lemma lim:
+  assumes J: "finite J" and X: "X \<in> sets (PiM J M)"
+  shows "emeasure PF.lim (PF.emb UNIV J X) = emeasure (CI J) X"
+proof (rule PF.emeasure_lim[OF J subset_UNIV X])
+  fix J X' assume J[simp]: "\<And>i. finite (J i)" and X'[measurable]: "\<And>i. X' i \<in> sets (Pi\<^sub>M (J i) M)"
+    and dec: "decseq (\<lambda>i. PF.emb UNIV (J i) (X' i))"
+  def X \<equiv> "\<lambda>n. (\<Inter>i\<in>{i. J i \<subseteq> {0..< n}}. PF.emb {0..<n} (J i) (X' i)) \<inter> space (PiM {0..<n} M)"
+
+  have sets_X[measurable]: "X n \<in> sets (PiM {0..<n} M)" for n
+    by (cases "{i. J i \<subseteq> {0..< n}} = {}")
+       (simp_all add: X_def, auto intro!: sets.countable_INT' sets.Int)
+  
+  have dec_X: "n \<le> m \<Longrightarrow> X m \<subseteq> PF.emb {0..<m} {0..<n} (X n)" for n m
+    unfolding X_def using ivl_subset[of 0 n 0 m]
+    by (cases "{i. J i \<subseteq> {0..< n}} = {}")
+       (auto simp add: prod_emb_Int prod_emb_PiE space_PiM simp del: ivl_subset)
+
+  have dec_X': "PF.emb {0..<n} (J j) (X' j) \<subseteq> PF.emb {0..<n} (J i) (X' i)"
+    if [simp]: "J i \<subseteq> {0..<n}" "J j \<subseteq> {0..<n}" "i \<le> j" for n i j
+    by (rule PF.emb_preserve_mono[of "{0..<n}" UNIV]) (auto del: subsetI intro: dec[THEN antimonoD])
+
+  assume "0 < (INF i. CI (J i) (X' i))"
+  also have "\<dots> \<le> (INF i. C 0 i (\<lambda>x. undefined) (X i))"
+  proof (intro INF_greatest)
+    fix n
+    interpret C!: prob_space "C 0 n (\<lambda>x. undefined)"
+      by (rule prob_space_C) simp
+    show "(INF i. CI (J i) (X' i)) \<le> C 0 n (\<lambda>x. undefined) (X n)"
+    proof cases
+      assume "{i. J i \<subseteq> {0..< n}} = {}" with C.emeasure_space_1  show ?thesis
+        by (auto simp add: X_def space_C intro!: INF_lower2[of 0] prob_space.measure_le_1 PF.prob_space_P)
+    next
+      assume *: "{i. J i \<subseteq> {0..< n}} \<noteq> {}"
+      have "(INF i. CI (J i) (X' i)) \<le>
+          (INF i:{i. J i \<subseteq> {0..<n}}. C 0 n (\<lambda>_. undefined) (PF.emb {0..<n} (J i) (X' i)))"
+        by (intro INF_superset_mono) (auto simp: emeasure_CI)
+      also have "\<dots> = C 0 n (\<lambda>_. undefined) (\<Inter>i\<in>{i. J i \<subseteq> {0..<n}}. (PF.emb {0..<n} (J i) (X' i)))"
+        using * by (intro emeasure_INT_decseq_subset[symmetric]) (auto intro!: dec_X' del: subsetI simp: sets_C)
+      also have "\<dots> = C 0 n (\<lambda>_. undefined) (X n)"
+        using * by (auto simp add: X_def INT_extend_simps)
+      finally show "(INF i. CI (J i) (X' i)) \<le> C 0 n (\<lambda>_. undefined) (X n)" .
+    qed
   qed
+  finally have pos: "0 < (INF i. C 0 i (\<lambda>x. undefined) (X i))" .
+  from less_INF_D[OF this, of 0] have "X 0 \<noteq> {}"
+    by auto
+
+  { fix \<omega> n assume \<omega>: "\<omega> \<in> space (PiM {0..<n} M)"
+    let ?C = "\<lambda>i. emeasure (C n i \<omega>) (X (n + i))"
+    let ?C' = "\<lambda>i x. emeasure (C (Suc n) i (\<omega>(n:=x))) (X (Suc n + i))"
+    have M: "\<And>i. ?C' i \<in> borel_measurable (P n \<omega>)"
+      using \<omega>[measurable, simp] measurable_fun_upd[where J="{0..<n}"] by measurable auto
+
+    assume "0 < (INF i. ?C i)"
+    also have "\<dots> \<le> (INF i. emeasure (C n (1 + i) \<omega>) (X (n + (1 + i))))"
+      by (intro INF_greatest INF_lower) auto
+    also have "\<dots> = (INF i. \<integral>\<^sup>+x. ?C' i x \<partial>P n \<omega>)"
+      using \<omega> measurable_C[of "Suc n"]
+      apply (intro INF_cong refl)
+      apply (subst split_C[symmetric, OF \<omega>])
+      apply (subst emeasure_bind[OF _ _ sets_X])
+      apply (simp_all del: C.simps add: space_C)
+      apply measurable
+      apply simp
+      apply (simp add: bind_return[OF measurable_eP] nn_integral_eP)
+      done
+    also have "\<dots> = (\<integral>\<^sup>+x. (INF i. ?C' i x) \<partial>P n \<omega>)"
+    proof (rule nn_integral_monotone_convergence_INF[symmetric])
+      have "(\<integral>\<^sup>+x. ?C' 0 x \<partial>P n \<omega>) \<le> (\<integral>\<^sup>+x. 1 \<partial>P n \<omega>)"
+        by (intro nn_integral_mono) (auto split: split_indicator)
+      also have "\<dots> < \<infinity>"
+        using prob_space_P[OF \<omega>, THEN prob_space.emeasure_space_1] by simp
+      finally show "(\<integral>\<^sup>+x. ?C' 0 x \<partial>P n \<omega>) < \<infinity>" .
+    next
+      fix i j :: nat and x assume "i \<le> j" "x \<in> space (P n \<omega>)"
+      with \<omega> have \<omega>': "\<omega>(n := x) \<in> space (PiM {0..<Suc n} M)"
+        by (auto simp: space_P[OF \<omega>] space_PiM PiE_iff extensional_def)
+      show "?C' j x \<le> ?C' i x"
+        using \<open>i \<le> j\<close> by (subst emeasure_C[symmetric, of i]) (auto intro!: emeasure_mono dec_X del: subsetI simp: sets_C space_P \<omega> \<omega>')
+    qed fact
+    finally have "(\<integral>\<^sup>+x. (INF i. ?C' i x) \<partial>P n \<omega>) \<noteq> 0"
+      by simp
+    then have "\<exists>\<^sub>F x in ae_filter (P n \<omega>). 0 < (INF i. ?C' i x)"
+       using M by (subst (asm) nn_integral_0_iff_AE)
+         (auto intro!: borel_measurable_INF simp: Filter.not_eventually not_le)
+    then have "\<exists>\<^sub>F x in ae_filter (P n \<omega>). x \<in> space (M n) \<and> 0 < (INF i. ?C' i x)"
+      by (rule frequently_mp[rotated]) (auto simp: space_P \<omega>)
+    then obtain x where "x \<in> space (M n)" "0 < (INF i. ?C' i x)"
+      by (auto dest: frequently_ex)
+    from this(2)[THEN less_INF_D, of 0] this(2)
+    have "\<exists>x. \<omega>(n := x) \<in> X (Suc n) \<and> 0 < (INF i. ?C' i x)"
+      by (intro exI[of _ x]) (simp split: split_indicator_asm) }
+  note step = this
+
+  let ?\<omega> = "\<lambda>\<omega> n x. (restrict \<omega> {0..<n})(n := x)"
+  let ?L = "\<lambda>\<omega> n r. INF i. emeasure (C (Suc n) i (?\<omega> \<omega> n r)) (X (Suc n + i))"
+  have *: "(\<And>i. i < n \<Longrightarrow> ?\<omega> \<omega> i (\<omega> i) \<in> X (Suc i)) \<Longrightarrow>
+    restrict \<omega> {0..<n} \<in> space (Pi\<^sub>M {0..<n} M)" for \<omega> n
+    using sets.sets_into_space[OF sets_X, of n]
+    by (cases n) (auto simp: atLeastLessThanSuc restrict_def[of _ "{}"])
+  have "\<exists>\<omega>. \<forall>n. ?\<omega> \<omega> n (\<omega> n) \<in> X (Suc n) \<and> 0 < ?L \<omega> n (\<omega> n)"
+  proof (rule dependent_wellorder_choice)
+    fix n \<omega> assume IH: "\<And>i. i < n \<Longrightarrow> ?\<omega> \<omega> i (\<omega> i) \<in> X (Suc i) \<and> 0 < ?L \<omega> i (\<omega> i)"
+    show "\<exists>r. ?\<omega> \<omega> n r \<in> X (Suc n) \<and> 0 < ?L \<omega> n r"
+    proof (rule step)
+      show "restrict \<omega> {0..<n} \<in> space (Pi\<^sub>M {0..<n} M)"
+        using IH[THEN conjunct1] by (rule *)
+      show "0 < (INF i. emeasure (C n i (restrict \<omega> {0..<n})) (X (n + i)))"
+      proof (cases n)
+        case 0 with pos show ?thesis
+          by (simp add: CI_def restrict_def)
+      next
+        case (Suc i) then show ?thesis
+          using IH[of i, THEN conjunct2] by (simp add: atLeastLessThanSuc)
+      qed
+    qed
+  qed (simp cong: restrict_cong)
+  then obtain \<omega> where \<omega>: "\<And>n. ?\<omega> \<omega> n (\<omega> n) \<in> X (Suc n)"
+    by auto
+  from this[THEN *] have \<omega>_space: "\<omega> \<in> space (PiM UNIV M)"
+    by (auto simp: space_PiM PiE_iff Ball_def)
+  have *: "\<omega> \<in> PF.emb UNIV {0..<n} (X n)" for n
+  proof (cases n)
+    case 0 with \<omega>_space \<open>X 0 \<noteq> {}\<close> sets.sets_into_space[OF sets_X, of 0] show ?thesis
+      by (auto simp add: space_PiM prod_emb_def restrict_def PiE_iff)
+  next
+    case (Suc i) then show ?thesis
+      using \<omega>[of i] \<omega>_space by (auto simp: prod_emb_def space_PiM PiE_iff atLeastLessThanSuc)
+  qed
+  have **: "{i. J i \<subseteq> {0..<up_to (J n)}} \<noteq> {}" for n
+    by (auto intro!: exI[of _ n] up_to J)
+  have "\<omega> \<in> PF.emb UNIV (J n) (X' n)" for n
+    using *[of "up_to (J n)"] up_to[of "J n"] by (simp add: X_def prod_emb_Int prod_emb_INT[OF **])
+  then show "(\<Inter>i. PF.emb UNIV (J i) (X' i)) \<noteq> {}"
+    by auto
+qed
+
+lemma distr_lim: assumes J[simp]: "finite J" shows "distr PF.lim (PiM J M) (\<lambda>x. restrict x J) = CI J"
+  apply (rule measure_eqI)
+  apply (simp add: CI_def)
+  apply (simp add: emeasure_distr measurable_cong_sets[OF PF.sets_lim] lim[symmetric] prod_emb_def space_PiM)
+  done
+
+end
+
+lemma (in product_prob_space) emeasure_lim_emb:
+  assumes *: "finite J" "J \<subseteq> I" "X \<in> sets (PiM J M)"
+  shows "emeasure lim (emb I J X) = emeasure (Pi\<^sub>M J M) X"
+proof (rule emeasure_lim[OF *], goal_cases)
+  case (1 J X)
+  
+  have "\<exists>Q. (\<forall>i. sets Q = PiM (\<Union>i. J i) M \<and> distr Q (PiM (J i) M) (\<lambda>x. restrict x (J i)) = Pi\<^sub>M (J i) M)"
+  proof cases
+    assume "finite (\<Union>i. J i)"
+    then have "distr (PiM (\<Union>i. J i) M) (Pi\<^sub>M (J i) M) (\<lambda>x. restrict x (J i)) = Pi\<^sub>M (J i) M" for i
+      by (intro distr_restrict[symmetric]) auto
+    then show ?thesis
+      by auto
+  next
+    assume inf: "infinite (\<Union>i. J i)"
+    moreover have count: "countable (\<Union>i. J i)"
+      using 1(3) by (auto intro: countable_finite)
+    def f \<equiv> "from_nat_into (\<Union>i. J i)" and t \<equiv> "to_nat_on (\<Union>i. J i)"
+    have ft[simp]: "x \<in> J i \<Longrightarrow> f (t x) = x" for x i
+      unfolding f_def t_def using inf count by (intro from_nat_into_to_nat_on) auto
+    have tf[simp]: "t (f i) = i" for i
+      unfolding t_def f_def by (intro to_nat_on_from_nat_into_infinite inf count)
+    have inj_t: "inj_on t (\<Union>i. J i)"
+      using count by (auto simp: t_def)
+    then have inj_t_J: "inj_on t (J i)" for i
+      by (rule subset_inj_on) auto
+    interpret IT!: Ionescu_Tulcea "\<lambda>i \<omega>. M (f i)" "\<lambda>i. M (f i)"
+      by standard auto
+    interpret Mf!: product_prob_space "\<lambda>x. M (f x)" UNIV
+      by standard
+    have C_eq_PiM: "IT.C 0 n (\<lambda>_. undefined) = PiM {0..<n} (\<lambda>x. M (f x))" for n
+    proof (induction n)
+      case 0 then show ?case
+        by (auto simp: PiM_empty intro!: measure_eqI dest!: subset_singletonD)
+    next
+      case (Suc n) then show ?case
+        apply (auto intro!: measure_eqI simp: sets_bind[OF IT.sets_eP] emeasure_bind[OF _ IT.measurable_eP])
+        apply (auto simp: Mf.product_nn_integral_insert nn_integral_indicator[symmetric] atLeastLessThanSuc IT.emeasure_eP space_PiM
+                    split: split_indicator simp del: nn_integral_indicator intro!: nn_integral_cong)
+        done
+    qed
+    have CI_eq_PiM: "IT.CI X = PiM X (\<lambda>x. M (f x))" if X: "finite X" for X
+      by (auto simp: IT.up_to_less X  IT.CI_def C_eq_PiM intro!: Mf.distr_restrict[symmetric])
+
+    let ?Q = "distr IT.PF.lim (PiM (\<Union>i. J i) M) (\<lambda>\<omega>. \<lambda>x\<in>\<Union>i. J i. \<omega> (t x))"
+    { fix i
+      have "distr ?Q (Pi\<^sub>M (J i) M) (\<lambda>x. restrict x (J i)) = 
+        distr IT.PF.lim (Pi\<^sub>M (J i) M) ((\<lambda>\<omega>. \<lambda>n\<in>J i. \<omega> (t n)) \<circ> (\<lambda>\<omega>. restrict \<omega> (t`J i)))"
+      proof (subst distr_distr)
+        have "(\<lambda>\<omega>. \<omega> (t x)) \<in> measurable (Pi\<^sub>M UNIV (\<lambda>x. M (f x))) (M x)" if x: "x \<in> J i" for x i
+          using measurable_component_singleton[of "t x" "UNIV" "\<lambda>x. M (f x)"] unfolding ft[OF x] by simp
+        then show "(\<lambda>\<omega>. \<lambda>x\<in>\<Union>i. J i. \<omega> (t x)) \<in> measurable IT.PF.lim (Pi\<^sub>M (UNION UNIV J) M)"
+          by (auto intro!: measurable_restrict simp: measurable_cong_sets[OF IT.PF.sets_lim refl])
+      qed (auto intro!: distr_cong measurable_restrict measurable_component_singleton)
+      also have "\<dots> = distr (distr IT.PF.lim (PiM (t`J i) (\<lambda>x. M (f x))) (\<lambda>\<omega>. restrict \<omega> (t`J i))) (Pi\<^sub>M (J i) M) (\<lambda>\<omega>. \<lambda>n\<in>J i. \<omega> (t n))"
+      proof (intro distr_distr[symmetric])
+        have "(\<lambda>\<omega>. \<omega> (t x)) \<in> measurable (Pi\<^sub>M (t`J i) (\<lambda>x. M (f x))) (M x)" if x: "x \<in> J i" for x
+          using measurable_component_singleton[of "t x" "t`J i" "\<lambda>x. M (f x)"] x unfolding ft[OF x] by auto
+        then show "(\<lambda>\<omega>. \<lambda>n\<in>J i. \<omega> (t n)) \<in> measurable (Pi\<^sub>M (t ` J i) (\<lambda>x. M (f x))) (Pi\<^sub>M (J i) M)"
+          by (auto intro!: measurable_restrict)
+      qed (auto intro!: measurable_restrict simp: measurable_cong_sets[OF IT.PF.sets_lim refl])
+      also have "\<dots> = distr (PiM (t`J i) (\<lambda>x. M (f x))) (Pi\<^sub>M (J i) M) (\<lambda>\<omega>. \<lambda>n\<in>J i. \<omega> (t n))"
+        using \<open>finite (J i)\<close> by (subst IT.distr_lim) (auto simp: CI_eq_PiM)
+      also have "\<dots> = Pi\<^sub>M (J i) M"
+        using Mf.distr_reorder[of t "J i"] by (simp add: 1 inj_t_J cong: PiM_cong)
+      finally have "distr ?Q (Pi\<^sub>M (J i) M) (\<lambda>x. restrict x (J i)) = Pi\<^sub>M (J i) M" . }
+    then show "\<exists>Q. \<forall>i. sets Q = PiM (\<Union>i. J i) M \<and> distr Q (Pi\<^sub>M (J i) M) (\<lambda>x. restrict x (J i)) = Pi\<^sub>M (J i) M"
+      by (intro exI[of _ ?Q]) auto
+  qed
+  then obtain Q where sets_Q: "sets Q = PiM (\<Union>i. J i) M"
+    and Q: "\<And>i. distr Q (PiM (J i) M) (\<lambda>x. restrict x (J i)) = Pi\<^sub>M (J i) M" by blast
+
+  from 1 interpret Q: prob_space Q
+    by (intro prob_space_distrD[of "\<lambda>x. restrict x (J 0)" Q "PiM (J 0) M"])
+       (auto simp: Q measurable_cong_sets[OF sets_Q]
+                intro!: prob_space_P measurable_restrict measurable_component_singleton)
+
+  have "0 < (INF i. emeasure (Pi\<^sub>M (J i) M) (X i))" by fact
+  also have "\<dots> = (INF i. emeasure Q (emb (\<Union>i. J i) (J i) (X i)))"
+    by (simp add: emeasure_distr_restrict[OF _ sets_Q 1(4), symmetric] SUP_upper Q)
+  also have "\<dots> = emeasure Q (\<Inter>i. emb (\<Union>i. J i) (J i) (X i))"
+  proof (rule INF_emeasure_decseq)
+    from 1 show "decseq (\<lambda>n. emb (\<Union>i. J i) (J n) (X n))"
+      by (intro antimonoI emb_preserve_mono[where X="emb (\<Union>i. J i) (J n) (X n)" and L=I and J="\<Union>i. J i" for n]
+         measurable_prod_emb)
+         (auto simp: SUP_least SUP_upper antimono_def)
+  qed (insert 1, auto simp: sets_Q)
+  finally have "(\<Inter>i. emb (\<Union>i. J i) (J i) (X i)) \<noteq> {}"
+    by auto
+  moreover have "(\<Inter>i. emb I (J i) (X i)) = {} \<Longrightarrow> (\<Inter>i. emb (\<Union>i. J i) (J i) (X i)) = {}"
+    using 1 by (intro emb_injective[of "\<Union>i. J i" I _ "{}"] sets.countable_INT) (auto simp: SUP_least SUP_upper)
+  ultimately show ?case by auto
 qed
 
 end
--- a/src/HOL/Probability/Projective_Limit.thy	Wed Oct 07 15:31:59 2015 +0200
+++ b/src/HOL/Probability/Projective_Limit.thy	Wed Oct 07 17:11:16 2015 +0200
@@ -101,440 +101,361 @@
   for I::"'i set" and P
 begin
 
-abbreviation "lim\<^sub>B \<equiv> (\<lambda>J P. limP J (\<lambda>_. borel) P)"
+lemma emeasure_lim_emb:
+  assumes X: "J \<subseteq> I" "finite J" "X \<in> sets (\<Pi>\<^sub>M i\<in>J. borel)"
+  shows "lim (emb I J X) = P J X"
+proof (rule emeasure_lim)
+  write mu_G ("\<mu>G")
+  interpret generator: algebra "space (PiM I (\<lambda>i. borel))" generator
+    by (rule algebra_generator)
+
+  fix J and B :: "nat \<Rightarrow> ('i \<Rightarrow> 'a) set"
+  assume J: "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. B n \<in> sets (\<Pi>\<^sub>M i\<in>J n. borel)" "incseq J"
+    and B: "decseq (\<lambda>n. emb I (J n) (B n))"
+    and "0 < (INF i. P (J i) (B i))" (is "0 < ?a")
+  moreover have "?a \<le> 1"
+    using J by (auto intro!: INF_lower2[of 0] prob_space_P[THEN prob_space.measure_le_1])
+  ultimately obtain r where r: "?a = ereal r" "0 < r" "r \<le> 1"
+    by (cases "?a") auto
+  def Z \<equiv> "\<lambda>n. emb I (J n) (B n)"
+  have Z_mono: "n \<le> m \<Longrightarrow> Z m \<subseteq> Z n" for n m
+    unfolding Z_def using B[THEN antimonoD, of n m] .
+  have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
+    using \<open>incseq J\<close> by (force simp: incseq_def)
+  note [simp] = \<open>\<And>n. finite (J n)\<close>
+  interpret prob_space "P (J i)" for i using J prob_space_P by simp
+
+  have P_eq[simp]:
+      "sets (P (J i)) = sets (\<Pi>\<^sub>M i\<in>J i. borel)" "space (P (J i)) = space (\<Pi>\<^sub>M i\<in>J i. borel)" for i
+    using J by (auto simp: sets_P space_P)
+
+  have "Z i \<in> generator" for i
+    unfolding Z_def by (auto intro!: generator.intros J)
 
-lemma emeasure_limB_emb_not_empty:
-  assumes "I \<noteq> {}"
-  assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
-  shows "emeasure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = emeasure (lim\<^sub>B J P) (Pi\<^sub>E J B)"
-proof -
-  let ?\<Omega> = "\<Pi>\<^sub>E i\<in>I. space borel"
-  let ?G = generator
-  interpret G!: algebra ?\<Omega> generator by (intro  algebra_generator) fact
-  note mu_G_mono =
-    G.additive_increasing[OF positive_mu_G[OF `I \<noteq> {}`] additive_mu_G[OF `I \<noteq> {}`],
-      THEN increasingD]
-  write mu_G  ("\<mu>G")
-
-  have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>"
-  proof (rule G.caratheodory_empty_continuous[OF positive_mu_G additive_mu_G,
-      OF `I \<noteq> {}`, OF `I \<noteq> {}`])
-    fix A assume "A \<in> ?G"
-    with generatorE guess J X . note JX = this
-    interpret prob_space "P J" using proj_prob_space[OF `finite J`] .
-    show "\<mu>G A \<noteq> \<infinity>" using JX by (simp add: limP_finite)
-  next
-    fix Z assume Z: "range Z \<subseteq> ?G" "decseq Z" "(\<Inter>i. Z i) = {}"
-    then have "decseq (\<lambda>i. \<mu>G (Z i))"
-      by (auto intro!: mu_G_mono simp: decseq_def)
-    moreover
-    have "(INF i. \<mu>G (Z i)) = 0"
+  have countable_UN_J: "countable (\<Union>n. J n)" by (simp add: countable_finite)
+  def Utn \<equiv> "to_nat_on (\<Union>n. J n)"
+  interpret function_to_finmap "J n" Utn "from_nat_into (\<Union>n. J n)" for n
+    by unfold_locales (auto simp: Utn_def intro: from_nat_into_to_nat_on[OF countable_UN_J])
+  have inj_on_Utn: "inj_on Utn (\<Union>n. J n)"
+    unfolding Utn_def using countable_UN_J by (rule inj_on_to_nat_on)
+  hence inj_on_Utn_J: "\<And>n. inj_on Utn (J n)" by (rule subset_inj_on) auto
+  def P' \<equiv> "\<lambda>n. mapmeasure n (P (J n)) (\<lambda>_. borel)"
+  interpret P': prob_space "P' n" for n
+    unfolding P'_def mapmeasure_def using J
+    by (auto intro!: prob_space_distr fm_measurable simp: measurable_cong_sets[OF sets_P])
+  
+  let ?SUP = "\<lambda>n. SUP K : {K. K \<subseteq> fm n ` (B n) \<and> compact K}. emeasure (P' n) K"
+  { fix n
+    have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))"
+      using J by (auto simp: P'_def mapmeasure_PiM space_P sets_P)
+    also
+    have "\<dots> = ?SUP n"
+    proof (rule inner_regular)
+      show "sets (P' n) = sets borel" by (simp add: borel_eq_PiF_borel P'_def)
+    next
+      show "fm n ` B n \<in> sets borel"
+        unfolding borel_eq_PiF_borel by (auto simp: P'_def fm_image_measurable_finite sets_P J(3))
+    qed simp
+    finally have "emeasure (P (J n)) (B n) = ?SUP n" "?SUP n \<noteq> \<infinity>" "?SUP n \<noteq> - \<infinity>" by auto
+  } note R = this
+  have "\<forall>n. \<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> 2 powr (-n) * ?a \<and> compact K \<and> K \<subseteq> fm n ` B n"
+  proof
+    fix n show "\<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> ereal (2 powr - real n) * ?a \<and>
+        compact K \<and> K \<subseteq> fm n ` B n"
+      unfolding R[of n]
     proof (rule ccontr)
-      assume "(INF i. \<mu>G (Z i)) \<noteq> 0" (is "?a \<noteq> 0")
-      moreover have "0 \<le> ?a"
-        using Z positive_mu_G[OF `I \<noteq> {}`] by (auto intro!: INF_greatest simp: positive_def)
-      ultimately have "0 < ?a" by auto
-      hence "?a \<noteq> -\<infinity>" by auto
-      have "\<forall>n. \<exists>J B. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> B \<in> sets (Pi\<^sub>M J (\<lambda>_. borel)) \<and>
-        Z n = emb I J B \<and> \<mu>G (Z n) = emeasure (lim\<^sub>B J P) B"
-        using Z by (intro allI generator_Ex) auto
-      then obtain J' B' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I"
-          "\<And>n. B' n \<in> sets (\<Pi>\<^sub>M i\<in>J' n. borel)"
-        and Z_emb: "\<And>n. Z n = emb I (J' n) (B' n)"
-        unfolding choice_iff by blast
-      moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"
-      moreover def B \<equiv> "\<lambda>n. emb (J n) (J' n) (B' n)"
-      ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I"
-        "\<And>n. B n \<in> sets (\<Pi>\<^sub>M i\<in>J n. borel)"
-        by auto
-      have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
-        unfolding J_def by force
-      have "\<forall>n. \<exists>j. j \<in> J n" using J by blast
-      then obtain j where j: "\<And>n. j n \<in> J n"
-        unfolding choice_iff by blast
-      note [simp] = `\<And>n. finite (J n)`
-      from J  Z_emb have Z_eq: "\<And>n. Z n = emb I (J n) (B n)" "\<And>n. Z n \<in> ?G"
-        unfolding J_def B_def by (subst prod_emb_trans) (insert Z, auto)
-      interpret prob_space "P (J i)" for i using proj_prob_space by simp
-      have "?a \<le> \<mu>G (Z 0)" by (auto intro: INF_lower)
-      also have "\<dots> < \<infinity>" using J by (auto simp: Z_eq mu_G_eq limP_finite proj_sets)
-      finally have "?a \<noteq> \<infinity>" by simp
-      have "\<And>n. \<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>" unfolding Z_eq using J J_mono
-        by (subst mu_G_eq) (auto simp: limP_finite proj_sets mu_G_eq)
+      assume H: "\<not> (\<exists>K'. ?SUP n - emeasure (P' n) K' \<le> ereal (2 powr - real n)  * ?a \<and>
+        compact K' \<and> K' \<subseteq> fm n ` B n)"
+      have "?SUP n \<le> ?SUP n - 2 powr (-n) * ?a"
+      proof (intro SUP_least)
+        fix K
+        assume "K \<in> {K. K \<subseteq> fm n ` B n \<and> compact K}"
+        with H have "\<not> ?SUP n - emeasure (P' n) K \<le> 2 powr (-n) * ?a"
+          by auto
+        hence "?SUP n - emeasure (P' n) K > 2 powr (-n) * ?a"
+          unfolding not_less[symmetric] by simp
+        hence "?SUP n - 2 powr (-n) * ?a > emeasure (P' n) K"
+          using `0 < ?a` by (auto simp add: ereal_less_minus_iff ac_simps)
+        thus "?SUP n - 2 powr (-n) * ?a \<ge> emeasure (P' n) K" by simp
+      qed
+      hence "?SUP n + 0 \<le> ?SUP n - (2 powr (-n) * ?a)" using `0 < ?a` by simp
+      hence "?SUP n + 0 \<le> ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def .
+      hence "0 \<le> - (2 powr (-n) * ?a)"
+        using `?SUP n \<noteq> \<infinity>` `?SUP n \<noteq> - \<infinity>`
+        by (subst (asm) ereal_add_le_add_iff) (auto simp:)
+      moreover have "ereal (2 powr - real n) * ?a > 0" using `0 < ?a`
+        by (auto simp: ereal_zero_less_0_iff)
+      ultimately show False by simp
+    qed
+  qed
+  then obtain K' where K':
+    "\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ereal (2 powr - real n) * ?a"
+    "\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n"
+    unfolding choice_iff by blast
+  def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
+  have K_sets: "\<And>n. K n \<in> sets (Pi\<^sub>M (J n) (\<lambda>_. borel))"
+    unfolding K_def
+    using compact_imp_closed[OF `compact (K' _)`]
+    by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"])
+       (auto simp: borel_eq_PiF_borel[symmetric])
+  have K_B: "\<And>n. K n \<subseteq> B n"
+  proof
+    fix x n assume "x \<in> K n"
+    then have fm_in: "fm n x \<in> fm n ` B n"
+      using K' by (force simp: K_def)
+    show "x \<in> B n"
+      using `x \<in> K n` K_sets sets.sets_into_space J(1,2,3)[of n]
+      by (intro inj_on_image_mem_iff[OF inj_on_fm _ fm_in, of "\<lambda>_. borel"]) auto
+  qed
+  def Z' \<equiv> "\<lambda>n. emb I (J n) (K n)"
+  have Z': "\<And>n. Z' n \<subseteq> Z n"
+    unfolding Z'_def Z_def
+  proof (rule prod_emb_mono, safe)
+    fix n x assume "x \<in> K n"
+    hence "fm n x \<in> K' n" "x \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
+      by (simp_all add: K_def space_P)
+    note this(1)
+    also have "K' n \<subseteq> fm n ` B n" by (simp add: K')
+    finally have "fm n x \<in> fm n ` B n" .
+    thus "x \<in> B n"
+    proof safe
+      fix y assume y: "y \<in> B n"
+      hence "y \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" using J sets.sets_into_space[of "B n" "P (J n)"]
+        by (auto simp add: space_P sets_P)
+      assume "fm n x = fm n y"
+      note inj_onD[OF inj_on_fm[OF space_borel],
+        OF `fm n x = fm n y` `x \<in> space _` `y \<in> space _`]
+      with y show "x \<in> B n" by simp
+    qed
+  qed
+  have "\<And>n. Z' n \<in> generator" using J K'(2) unfolding Z'_def
+    by (auto intro!: generator.intros measurable_sets[OF fm_measurable[of _ "Collect finite"]]
+             simp: K_def borel_eq_PiF_borel[symmetric] compact_imp_closed)
+  def Y \<equiv> "\<lambda>n. \<Inter>i\<in>{1..n}. Z' i"
+  hence "\<And>n k. Y (n + k) \<subseteq> Y n" by (induct_tac k) (auto simp: Y_def)
+  hence Y_mono: "\<And>n m. n \<le> m \<Longrightarrow> Y m \<subseteq> Y n" by (auto simp: le_iff_add)
+  have Y_Z': "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z' n" by (auto simp: Y_def)
+  hence Y_Z: "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z n" using Z' by auto
 
-      have countable_UN_J: "countable (\<Union>n. J n)" by (simp add: countable_finite)
-      def Utn \<equiv> "to_nat_on (\<Union>n. J n)"
-      interpret function_to_finmap "J n" Utn "from_nat_into (\<Union>n. J n)" for n
-        by unfold_locales (auto simp: Utn_def intro: from_nat_into_to_nat_on[OF countable_UN_J])
-      have inj_on_Utn: "inj_on Utn (\<Union>n. J n)"
-        unfolding Utn_def using countable_UN_J by (rule inj_on_to_nat_on)
-      hence inj_on_Utn_J: "\<And>n. inj_on Utn (J n)" by (rule subset_inj_on) auto
-      def P' \<equiv> "\<lambda>n. mapmeasure n (P (J n)) (\<lambda>_. borel)"
-      let ?SUP = "\<lambda>n. SUP K : {K. K \<subseteq> fm n ` (B n) \<and> compact K}. emeasure (P' n) K"
-      {
-        fix n
-        interpret finite_measure "P (J n)" by unfold_locales
-        have "emeasure (P (J n)) (B n) = emeasure (P' n) (fm n ` (B n))"
-          using J by (auto simp: P'_def mapmeasure_PiM proj_space proj_sets)
-        also
-        have "\<dots> = ?SUP n"
-        proof (rule inner_regular)
-          show "emeasure (P' n) (space (P' n)) \<noteq> \<infinity>"
-            unfolding P'_def
-            by (auto simp: P'_def mapmeasure_PiF fm_measurable proj_space proj_sets)
-          show "sets (P' n) = sets borel" by (simp add: borel_eq_PiF_borel P'_def)
-        next
-          show "fm n ` B n \<in> sets borel"
-            unfolding borel_eq_PiF_borel
-            by (auto simp del: J(2) simp: P'_def fm_image_measurable_finite proj_sets J)
-        qed
-        finally
-        have "emeasure (P (J n)) (B n) = ?SUP n" "?SUP n \<noteq> \<infinity>" "?SUP n \<noteq> - \<infinity>" by auto
-      } note R = this
-      have "\<forall>n. \<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> 2 powr (-n) * ?a
-        \<and> compact K \<and> K \<subseteq> fm n ` B n"
-      proof
-        fix n
-        have "emeasure (P' n) (space (P' n)) \<noteq> \<infinity>"
-          by (simp add: mapmeasure_PiF P'_def proj_space proj_sets)
-        then interpret finite_measure "P' n" ..
-        show "\<exists>K. emeasure (P (J n)) (B n) - emeasure (P' n) K \<le> ereal (2 powr - real n) * ?a \<and>
-            compact K \<and> K \<subseteq> fm n ` B n"
-          unfolding R
-        proof (rule ccontr)
-          assume H: "\<not> (\<exists>K'. ?SUP n - emeasure (P' n) K' \<le> ereal (2 powr - real n)  * ?a \<and>
-            compact K' \<and> K' \<subseteq> fm n ` B n)"
-          have "?SUP n \<le> ?SUP n - 2 powr (-n) * ?a"
-          proof (intro SUP_least)
-            fix K
-            assume "K \<in> {K. K \<subseteq> fm n ` B n \<and> compact K}"
-            with H have "\<not> ?SUP n - emeasure (P' n) K \<le> 2 powr (-n) * ?a"
-              by auto
-            hence "?SUP n - emeasure (P' n) K > 2 powr (-n) * ?a"
-              unfolding not_less[symmetric] by simp
-            hence "?SUP n - 2 powr (-n) * ?a > emeasure (P' n) K"
-              using `0 < ?a` by (auto simp add: ereal_less_minus_iff ac_simps)
-            thus "?SUP n - 2 powr (-n) * ?a \<ge> emeasure (P' n) K" by simp
-          qed
-          hence "?SUP n + 0 \<le> ?SUP n - (2 powr (-n) * ?a)" using `0 < ?a` by simp
-          hence "?SUP n + 0 \<le> ?SUP n + - (2 powr (-n) * ?a)" unfolding minus_ereal_def .
-          hence "0 \<le> - (2 powr (-n) * ?a)"
-            using `?SUP _ \<noteq> \<infinity>` `?SUP _ \<noteq> - \<infinity>`
-            by (subst (asm) ereal_add_le_add_iff) (auto simp:)
-          moreover have "ereal (2 powr - real n) * ?a > 0" using `0 < ?a`
-            by (auto simp: ereal_zero_less_0_iff)
-          ultimately show False by simp
-        qed
-      qed
-      then obtain K' where K':
-        "\<And>n. emeasure (P (J n)) (B n) - emeasure (P' n) (K' n) \<le> ereal (2 powr - real n) * ?a"
-        "\<And>n. compact (K' n)" "\<And>n. K' n \<subseteq> fm n ` B n"
-        unfolding choice_iff by blast
-      def K \<equiv> "\<lambda>n. fm n -` K' n \<inter> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
-      have K_sets: "\<And>n. K n \<in> sets (Pi\<^sub>M (J n) (\<lambda>_. borel))"
-        unfolding K_def
-        using compact_imp_closed[OF `compact (K' _)`]
-        by (intro measurable_sets[OF fm_measurable, of _ "Collect finite"])
-           (auto simp: borel_eq_PiF_borel[symmetric])
-      have K_B: "\<And>n. K n \<subseteq> B n"
-      proof
-        fix x n
-        assume "x \<in> K n" hence fm_in: "fm n x \<in> fm n ` B n"
-          using K' by (force simp: K_def)
-        show "x \<in> B n"
-          using `x \<in> K n` K_sets sets.sets_into_space J[of n]
-          by (intro inj_on_image_mem_iff[OF inj_on_fm _ fm_in, of "\<lambda>_. borel"]) auto
-      qed
-      def Z' \<equiv> "\<lambda>n. emb I (J n) (K n)"
-      have Z': "\<And>n. Z' n \<subseteq> Z n"
-        unfolding Z_eq unfolding Z'_def
-      proof (rule prod_emb_mono, safe)
-        fix n x assume "x \<in> K n"
-        hence "fm n x \<in> K' n" "x \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))"
-          by (simp_all add: K_def proj_space)
-        note this(1)
-        also have "K' n \<subseteq> fm n ` B n" by (simp add: K')
-        finally have "fm n x \<in> fm n ` B n" .
-        thus "x \<in> B n"
-        proof safe
-          fix y assume y: "y \<in> B n"
-          hence "y \<in> space (Pi\<^sub>M (J n) (\<lambda>_. borel))" using J sets.sets_into_space[of "B n" "P (J n)"]
-            by (auto simp add: proj_space proj_sets)
-          assume "fm n x = fm n y"
-          note inj_onD[OF inj_on_fm[OF space_borel],
-            OF `fm n x = fm n y` `x \<in> space _` `y \<in> space _`]
-          with y show "x \<in> B n" by simp
+  have Y_notempty: "\<And>n. n \<ge> 1 \<Longrightarrow> (Y n) \<noteq> {}"
+  proof -
+    fix n::nat assume "n \<ge> 1" hence "Y n \<subseteq> Z n" by fact
+    have "Y n = (\<Inter>i\<in>{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono
+      by (auto simp: Y_def Z'_def)
+    also have "\<dots> = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. emb (J n) (J i) (K i))"
+      using `n \<ge> 1`
+      by (subst prod_emb_INT) auto
+    finally
+    have Y_emb:
+      "Y n = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))" .
+    hence "Y n \<in> generator" using J J_mono K_sets `n \<ge> 1`
+      by (auto simp del: prod_emb_INT intro!: generator.intros)
+    have *: "\<mu>G (Z n) = P (J n) (B n)"
+      unfolding Z_def using J by (intro mu_G_spec) auto
+    then have "\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>" by auto
+    note *
+    moreover have *: "\<mu>G (Y n) = P (J n) (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))"
+      unfolding Y_emb using J J_mono K_sets `n \<ge> 1` by (subst mu_G_spec) auto
+    then have "\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>" by auto
+    note *
+    moreover have "\<mu>G (Z n - Y n) =
+        P (J n) (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
+      unfolding Z_def Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets `n \<ge> 1`
+      by (subst mu_G_spec) (auto intro!: sets.Diff)
+    ultimately
+    have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)"
+      using J J_mono K_sets `n \<ge> 1`
+      by (simp only: emeasure_eq_measure Z_def)
+        (auto dest!: bspec[where x=n] intro!: measure_Diff[symmetric] set_mp[OF K_B]
+          simp: extensional_restrict emeasure_eq_measure prod_emb_iff sets_P space_P)
+    also have subs: "Z n - Y n \<subseteq> (\<Union>i\<in>{1..n}. (Z i - Z' i))"
+      using `n \<ge> 1` unfolding Y_def UN_extend_simps(7) by (intro UN_mono Diff_mono Z_mono order_refl) auto
+    have "Z n - Y n \<in> generator" "(\<Union>i\<in>{1..n}. (Z i - Z' i)) \<in> generator"
+      using `Z' _ \<in> generator` `Z _ \<in> generator` `Y _ \<in> generator` by auto
+    hence "\<mu>G (Z n - Y n) \<le> \<mu>G (\<Union>i\<in>{1..n}. (Z i - Z' i))"
+      using subs generator.additive_increasing[OF positive_mu_G additive_mu_G]
+      unfolding increasing_def by auto
+    also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using `Z _ \<in> generator` `Z' _ \<in> generator`
+      by (intro generator.subadditive[OF positive_mu_G additive_mu_G]) auto
+    also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
+    proof (rule setsum_mono)
+      fix i assume "i \<in> {1..n}" hence "i \<le> n" by simp
+      have "\<mu>G (Z i - Z' i) = \<mu>G (prod_emb I (\<lambda>_. borel) (J i) (B i - K i))"
+        unfolding Z'_def Z_def by simp
+      also have "\<dots> = P (J i) (B i - K i)"
+        using J K_sets by (subst mu_G_spec) auto
+      also have "\<dots> = P (J i) (B i) - P (J i) (K i)"
+        using K_sets J `K _ \<subseteq> B _` by (simp add: emeasure_Diff)
+      also have "\<dots> = P (J i) (B i) - P' i (K' i)"
+        unfolding K_def P'_def
+        by (auto simp: mapmeasure_PiF borel_eq_PiF_borel[symmetric]
+          compact_imp_closed[OF `compact (K' _)`] space_PiM PiE_def)
+      also have "\<dots> \<le> ereal (2 powr - real i) * ?a" using K'(1)[of i] .
+      finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" .
+    qed
+    also have "\<dots> = ereal ((\<Sum> i\<in>{1..n}. (2 powr -real i)) * real ?a)"
+      by (simp add: setsum_left_distrib r)
+    also have "\<dots> < ereal (1 * real ?a)" unfolding less_ereal.simps
+    proof (rule mult_strict_right_mono)
+      have "(\<Sum>i = 1..n. 2 powr - real i) = (\<Sum>i = 1..<Suc n. (1/2) ^ i)"
+        by (rule setsum.cong) (auto simp: powr_realpow powr_divide powr_minus_divide)  
+      also have "{1..<Suc n} = {..<Suc n} - {0}" by auto
+      also have "setsum (op ^ (1 / 2::real)) ({..<Suc n} - {0}) =
+        setsum (op ^ (1 / 2)) ({..<Suc n}) - 1" by (auto simp: setsum_diff1)
+      also have "\<dots> < 1" by (subst geometric_sum) auto
+      finally show "(\<Sum>i = 1..n. 2 powr - real i) < 1" .
+    qed (auto simp: r ereal_less_real_iff zero_ereal_def[symmetric])
+    also have "\<dots> = ?a" by (auto simp: r)
+    also have "\<dots> \<le> \<mu>G (Z n)"
+      using J by (auto intro: INF_lower simp: Z_def mu_G_spec)
+    finally have "\<mu>G (Z n) - \<mu>G (Y n) < \<mu>G (Z n)" .
+    hence R: "\<mu>G (Z n) < \<mu>G (Z n) + \<mu>G (Y n)"
+      using `\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>` by (simp add: ereal_minus_less)
+    have "0 \<le> (- \<mu>G (Z n)) + \<mu>G (Z n)" using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
+    also have "\<dots> < (- \<mu>G (Z n)) + (\<mu>G (Z n) + \<mu>G (Y n))"
+      apply (rule ereal_less_add[OF _ R]) using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
+    finally have "\<mu>G (Y n) > 0"
+      using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by (auto simp: ac_simps zero_ereal_def[symmetric])
+    thus "Y n \<noteq> {}" using positive_mu_G by (auto simp add: positive_def)
+  qed
+  hence "\<forall>n\<in>{1..}. \<exists>y. y \<in> Y n" by auto
+  then obtain y where y: "\<And>n. n \<ge> 1 \<Longrightarrow> y n \<in> Y n" unfolding bchoice_iff by force
+  {
+    fix t and n m::nat
+    assume "1 \<le> n" "n \<le> m" hence "1 \<le> m" by simp
+    from Y_mono[OF `m \<ge> n`] y[OF `1 \<le> m`] have "y m \<in> Y n" by auto
+    also have "\<dots> \<subseteq> Z' n" using Y_Z'[OF `1 \<le> n`] .
+    finally
+    have "fm n (restrict (y m) (J n)) \<in> K' n"
+      unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
+    moreover have "finmap_of (J n) (restrict (y m) (J n)) = finmap_of (J n) (y m)"
+      using J by (simp add: fm_def)
+    ultimately have "fm n (y m) \<in> K' n" by simp
+  } note fm_in_K' = this
+  interpret finmap_seqs_into_compact "\<lambda>n. K' (Suc n)" "\<lambda>k. fm (Suc k) (y (Suc k))" borel
+  proof
+    fix n show "compact (K' n)" by fact
+  next
+    fix n
+    from Y_mono[of n "Suc n"] y[of "Suc n"] have "y (Suc n) \<in> Y (Suc n)" by auto
+    also have "\<dots> \<subseteq> Z' (Suc n)" using Y_Z' by auto
+    finally
+    have "fm (Suc n) (restrict (y (Suc n)) (J (Suc n))) \<in> K' (Suc n)"
+      unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
+    thus "K' (Suc n) \<noteq> {}" by auto
+    fix k
+    assume "k \<in> K' (Suc n)"
+    with K'[of "Suc n"] sets.sets_into_space have "k \<in> fm (Suc n) ` B (Suc n)" by auto
+    then obtain b where "k = fm (Suc n) b" by auto
+    thus "domain k = domain (fm (Suc n) (y (Suc n)))"
+      by (simp_all add: fm_def)
+  next
+    fix t and n m::nat
+    assume "n \<le> m" hence "Suc n \<le> Suc m" by simp
+    assume "t \<in> domain (fm (Suc n) (y (Suc n)))"
+    then obtain j where j: "t = Utn j" "j \<in> J (Suc n)" by auto
+    hence "j \<in> J (Suc m)" using J_mono[OF `Suc n \<le> Suc m`] by auto
+    have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using `n \<le> m`
+      by (intro fm_in_K') simp_all
+    show "(fm (Suc m) (y (Suc m)))\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K' (Suc n)"
+      apply (rule image_eqI[OF _ img])
+      using `j \<in> J (Suc n)` `j \<in> J (Suc m)`
+      unfolding j by (subst proj_fm, auto)+
+  qed
+  have "\<forall>t. \<exists>z. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) ----> z"
+    using diagonal_tendsto ..
+  then obtain z where z:
+    "\<And>t. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) ----> z t"
+    unfolding choice_iff by blast
+  {
+    fix n :: nat assume "n \<ge> 1"
+    have "\<And>i. domain (fm n (y (Suc (diagseq i)))) = domain (finmap_of (Utn ` J n) z)"
+      by simp
+    moreover
+    {
+      fix t
+      assume t: "t \<in> domain (finmap_of (Utn ` J n) z)"
+      hence "t \<in> Utn ` J n" by simp
+      then obtain j where j: "t = Utn j" "j \<in> J n" by auto
+      have "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) ----> z t"
+        apply (subst (2) tendsto_iff, subst eventually_sequentially)
+      proof safe
+        fix e :: real assume "0 < e"
+        { fix i and x :: "'i \<Rightarrow> 'a" assume i: "i \<ge> n"
+          assume "t \<in> domain (fm n x)"
+          hence "t \<in> domain (fm i x)" using J_mono[OF `i \<ge> n`] by auto
+          with i have "(fm i x)\<^sub>F t = (fm n x)\<^sub>F t"
+            using j by (auto simp: proj_fm dest!: inj_onD[OF inj_on_Utn])
+        } note index_shift = this
+        have I: "\<And>i. i \<ge> n \<Longrightarrow> Suc (diagseq i) \<ge> n"
+          apply (rule le_SucI)
+          apply (rule order_trans) apply simp
+          apply (rule seq_suble[OF subseq_diagseq])
+          done
+        from z
+        have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e"
+          unfolding tendsto_iff eventually_sequentially using `0 < e` by auto
+        then obtain N where N: "\<And>i. i \<ge> N \<Longrightarrow>
+          dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" by auto
+        show "\<exists>N. \<forall>na\<ge>N. dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e "
+        proof (rule exI[where x="max N n"], safe)
+          fix na assume "max N n \<le> na"
+          hence  "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) =
+                  dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^sub>F t) (z t)" using t
+            by (subst index_shift[OF I]) auto
+          also have "\<dots> < e" using `max N n \<le> na` by (intro N) simp
+          finally show "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e" .
         qed
       qed
-      { fix n
-        have "Z' n \<in> ?G" using K' unfolding Z'_def
-          apply (intro generatorI'[OF J(1-3)])
-          unfolding K_def proj_space
-          apply (rule measurable_sets[OF fm_measurable[of _ "Collect finite"]])
-          apply (auto simp add: P'_def borel_eq_PiF_borel[symmetric] compact_imp_closed)
-          done
-      }
-      def Y \<equiv> "\<lambda>n. \<Inter>i\<in>{1..n}. Z' i"
-      hence "\<And>n k. Y (n + k) \<subseteq> Y n" by (induct_tac k) (auto simp: Y_def)
-      hence Y_mono: "\<And>n m. n \<le> m \<Longrightarrow> Y m \<subseteq> Y n" by (auto simp: le_iff_add)
-      have Y_Z': "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z' n" by (auto simp: Y_def)
-      hence Y_Z: "\<And>n. n \<ge> 1 \<Longrightarrow> Y n \<subseteq> Z n" using Z' by auto
-      have Y_notempty: "\<And>n. n \<ge> 1 \<Longrightarrow> (Y n) \<noteq> {}"
-      proof -
-        fix n::nat assume "n \<ge> 1" hence "Y n \<subseteq> Z n" by fact
-        have "Y n = (\<Inter>i\<in>{1..n}. emb I (J n) (emb (J n) (J i) (K i)))" using J J_mono
-          by (auto simp: Y_def Z'_def)
-        also have "\<dots> = prod_emb I (\<lambda>_. borel) (J n) (\<Inter>i\<in>{1..n}. emb (J n) (J i) (K i))"
-          using `n \<ge> 1`
-          by (subst prod_emb_INT) auto
-        finally
-        have Y_emb:
-          "Y n = prod_emb I (\<lambda>_. borel) (J n)
-            (\<Inter>i\<in>{1..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))" .
-        hence "Y n \<in> ?G" using J J_mono K_sets `n \<ge> 1` by (intro generatorI[OF _ _ _ _ Y_emb]) auto
-        hence "\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>" unfolding Y_emb using J J_mono K_sets `n \<ge> 1`
-          by (subst mu_G_eq) (auto simp: limP_finite proj_sets mu_G_eq)
-        interpret finite_measure "(limP (J n) (\<lambda>_. borel) P)"
-        proof
-          have "emeasure (limP (J n) (\<lambda>_. borel) P) (J n \<rightarrow>\<^sub>E space borel) \<noteq> \<infinity>"
-            using J by (subst emeasure_limP) auto
-          thus  "emeasure (limP (J n) (\<lambda>_. borel) P) (space (limP (J n) (\<lambda>_. borel) P)) \<noteq> \<infinity>"
-             by (simp add: space_PiM)
-        qed
-        have "\<mu>G (Z n) = limP (J n) (\<lambda>_. borel) P (B n)"
-          unfolding Z_eq using J by (auto simp: mu_G_eq)
-        moreover have "\<mu>G (Y n) =
-          limP (J n) (\<lambda>_. borel) P (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i))"
-          unfolding Y_emb using J J_mono K_sets `n \<ge> 1` by (subst mu_G_eq) auto
-        moreover have "\<mu>G (Z n - Y n) = limP (J n) (\<lambda>_. borel) P
-          (B n - (\<Inter>i\<in>{Suc 0..n}. prod_emb (J n) (\<lambda>_. borel) (J i) (K i)))"
-          unfolding Z_eq Y_emb prod_emb_Diff[symmetric] using J J_mono K_sets `n \<ge> 1`
-          by (subst mu_G_eq) (auto intro!: sets.Diff)
-        ultimately
-        have "\<mu>G (Z n) - \<mu>G (Y n) = \<mu>G (Z n - Y n)"
-          using J J_mono K_sets `n \<ge> 1`
-          by (simp only: emeasure_eq_measure)
-            (auto dest!: bspec[where x=n]
-            simp: extensional_restrict emeasure_eq_measure prod_emb_iff simp del: limP_finite
-            intro!: measure_Diff[symmetric] set_mp[OF K_B])
-        also have subs: "Z n - Y n \<subseteq> (\<Union>i\<in>{1..n}. (Z i - Z' i))" using Z' Z `n \<ge> 1`
-          unfolding Y_def by (force simp: decseq_def)
-        have "Z n - Y n \<in> ?G" "(\<Union>i\<in>{1..n}. (Z i - Z' i)) \<in> ?G"
-          using `Z' _ \<in> ?G` `Z _ \<in> ?G` `Y _ \<in> ?G` by auto
-        hence "\<mu>G (Z n - Y n) \<le> \<mu>G (\<Union>i\<in>{1..n}. (Z i - Z' i))"
-          using subs G.additive_increasing[OF positive_mu_G[OF `I \<noteq> {}`] additive_mu_G[OF `I \<noteq> {}`]]
-          unfolding increasing_def by auto
-        also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. \<mu>G (Z i - Z' i))" using `Z _ \<in> ?G` `Z' _ \<in> ?G`
-          by (intro G.subadditive[OF positive_mu_G additive_mu_G, OF `I \<noteq> {}` `I \<noteq> {}`]) auto
-        also have "\<dots> \<le> (\<Sum> i\<in>{1..n}. 2 powr -real i * ?a)"
-        proof (rule setsum_mono)
-          fix i assume "i \<in> {1..n}" hence "i \<le> n" by simp
-          have "\<mu>G (Z i - Z' i) = \<mu>G (prod_emb I (\<lambda>_. borel) (J i) (B i - K i))"
-            unfolding Z'_def Z_eq by simp
-          also have "\<dots> = P (J i) (B i - K i)"
-            using J K_sets by (subst mu_G_eq) auto
-          also have "\<dots> = P (J i) (B i) - P (J i) (K i)"
-            apply (subst emeasure_Diff) using K_sets J `K _ \<subseteq> B _` apply (auto simp: proj_sets)
-            done
-          also have "\<dots> = P (J i) (B i) - P' i (K' i)"
-            unfolding K_def P'_def
-            by (auto simp: mapmeasure_PiF proj_space proj_sets borel_eq_PiF_borel[symmetric]
-              compact_imp_closed[OF `compact (K' _)`] space_PiM PiE_def)
-          also have "\<dots> \<le> ereal (2 powr - real i) * ?a" using K'(1)[of i] .
-          finally show "\<mu>G (Z i - Z' i) \<le> (2 powr - real i) * ?a" .
-        qed
-        also have "\<dots> = (\<Sum> i\<in>{1..n}. ereal (2 powr -real i) * ereal(real ?a))"
-          using `?a \<noteq> \<infinity>` `?a \<noteq> - \<infinity>` by (subst ereal_real') auto
-        also have "\<dots> = ereal (\<Sum> i\<in>{1..n}. (2 powr -real i) * (real ?a))" by simp
-        also have "\<dots> = ereal ((\<Sum> i\<in>{1..n}. (2 powr -real i)) * real ?a)"
-          by (simp add: setsum_left_distrib)
-        also have "\<dots> < ereal (1 * real ?a)" unfolding less_ereal.simps
-        proof (rule mult_strict_right_mono)
-          have "(\<Sum>i = 1..n. 2 powr - real i) = (\<Sum>i = 1..<Suc n. (1/2) ^ i)"
-            by (rule setsum.cong) (auto simp: powr_realpow powr_divide powr_minus_divide)  
-          also have "{1..<Suc n} = {..<Suc n} - {0}" by auto
-          also have "setsum (op ^ (1 / 2::real)) ({..<Suc n} - {0}) =
-            setsum (op ^ (1 / 2)) ({..<Suc n}) - 1" by (auto simp: setsum_diff1)
-          also have "\<dots> < 1" by (subst geometric_sum) auto
-          finally show "(\<Sum>i = 1..n. 2 powr - real i) < 1" .
-        qed (auto simp:
-          `0 < ?a` `?a \<noteq> \<infinity>` `?a \<noteq> - \<infinity>` ereal_less_real_iff zero_ereal_def[symmetric])
-        also have "\<dots> = ?a" using `0 < ?a` `?a \<noteq> \<infinity>` by (auto simp: ereal_real')
-        also have "\<dots> \<le> \<mu>G (Z n)" by (auto intro: INF_lower)
-        finally have "\<mu>G (Z n) - \<mu>G (Y n) < \<mu>G (Z n)" .
-        hence R: "\<mu>G (Z n) < \<mu>G (Z n) + \<mu>G (Y n)"
-          using `\<bar>\<mu>G (Y n)\<bar> \<noteq> \<infinity>` by (simp add: ereal_minus_less)
-        have "0 \<le> (- \<mu>G (Z n)) + \<mu>G (Z n)" using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
-        also have "\<dots> < (- \<mu>G (Z n)) + (\<mu>G (Z n) + \<mu>G (Y n))"
-          apply (rule ereal_less_add[OF _ R]) using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by auto
-        finally have "\<mu>G (Y n) > 0"
-          using `\<bar>\<mu>G (Z n)\<bar> \<noteq> \<infinity>` by (auto simp: ac_simps zero_ereal_def[symmetric])
-        thus "Y n \<noteq> {}" using positive_mu_G `I \<noteq> {}` by (auto simp add: positive_def)
-      qed
-      hence "\<forall>n\<in>{1..}. \<exists>y. y \<in> Y n" by auto
-      then obtain y where y: "\<And>n. n \<ge> 1 \<Longrightarrow> y n \<in> Y n" unfolding bchoice_iff by force
-      {
-        fix t and n m::nat
-        assume "1 \<le> n" "n \<le> m" hence "1 \<le> m" by simp
-        from Y_mono[OF `m \<ge> n`] y[OF `1 \<le> m`] have "y m \<in> Y n" by auto
-        also have "\<dots> \<subseteq> Z' n" using Y_Z'[OF `1 \<le> n`] .
-        finally
-        have "fm n (restrict (y m) (J n)) \<in> K' n"
-          unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
-        moreover have "finmap_of (J n) (restrict (y m) (J n)) = finmap_of (J n) (y m)"
-          using J by (simp add: fm_def)
-        ultimately have "fm n (y m) \<in> K' n" by simp
-      } note fm_in_K' = this
-      interpret finmap_seqs_into_compact "\<lambda>n. K' (Suc n)" "\<lambda>k. fm (Suc k) (y (Suc k))" borel
-      proof
-        fix n show "compact (K' n)" by fact
-      next
-        fix n
-        from Y_mono[of n "Suc n"] y[of "Suc n"] have "y (Suc n) \<in> Y (Suc n)" by auto
-        also have "\<dots> \<subseteq> Z' (Suc n)" using Y_Z' by auto
-        finally
-        have "fm (Suc n) (restrict (y (Suc n)) (J (Suc n))) \<in> K' (Suc n)"
-          unfolding Z'_def K_def prod_emb_iff by (simp add: Z'_def K_def prod_emb_iff)
-        thus "K' (Suc n) \<noteq> {}" by auto
-        fix k
-        assume "k \<in> K' (Suc n)"
-        with K'[of "Suc n"] sets.sets_into_space have "k \<in> fm (Suc n) ` B (Suc n)" by auto
-        then obtain b where "k = fm (Suc n) b" by auto
-        thus "domain k = domain (fm (Suc n) (y (Suc n)))"
-          by (simp_all add: fm_def)
-      next
-        fix t and n m::nat
-        assume "n \<le> m" hence "Suc n \<le> Suc m" by simp
-        assume "t \<in> domain (fm (Suc n) (y (Suc n)))"
-        then obtain j where j: "t = Utn j" "j \<in> J (Suc n)" by auto
-        hence "j \<in> J (Suc m)" using J_mono[OF `Suc n \<le> Suc m`] by auto
-        have img: "fm (Suc n) (y (Suc m)) \<in> K' (Suc n)" using `n \<le> m`
-          by (intro fm_in_K') simp_all
-        show "(fm (Suc m) (y (Suc m)))\<^sub>F t \<in> (\<lambda>k. (k)\<^sub>F t) ` K' (Suc n)"
-          apply (rule image_eqI[OF _ img])
-          using `j \<in> J (Suc n)` `j \<in> J (Suc m)`
-          unfolding j by (subst proj_fm, auto)+
-      qed
-      have "\<forall>t. \<exists>z. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) ----> z"
-        using diagonal_tendsto ..
-      then obtain z where z:
-        "\<And>t. (\<lambda>i. (fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) ----> z t"
-        unfolding choice_iff by blast
-      {
-        fix n :: nat assume "n \<ge> 1"
-        have "\<And>i. domain (fm n (y (Suc (diagseq i)))) = domain (finmap_of (Utn ` J n) z)"
-          by simp
-        moreover
-        {
-          fix t
-          assume t: "t \<in> domain (finmap_of (Utn ` J n) z)"
-          hence "t \<in> Utn ` J n" by simp
-          then obtain j where j: "t = Utn j" "j \<in> J n" by auto
-          have "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) ----> z t"
-            apply (subst (2) tendsto_iff, subst eventually_sequentially)
-          proof safe
-            fix e :: real assume "0 < e"
-            { fix i x
-              assume i: "i \<ge> n"
-              assume "t \<in> domain (fm n x)"
-              hence "t \<in> domain (fm i x)" using J_mono[OF `i \<ge> n`] by auto
-              with i have "(fm i x)\<^sub>F t = (fm n x)\<^sub>F t"
-                using j by (auto simp: proj_fm dest!: inj_onD[OF inj_on_Utn])
-            } note index_shift = this
-            have I: "\<And>i. i \<ge> n \<Longrightarrow> Suc (diagseq i) \<ge> n"
-              apply (rule le_SucI)
-              apply (rule order_trans) apply simp
-              apply (rule seq_suble[OF subseq_diagseq])
-              done
-            from z
-            have "\<exists>N. \<forall>i\<ge>N. dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e"
-              unfolding tendsto_iff eventually_sequentially using `0 < e` by auto
-            then obtain N where N: "\<And>i. i \<ge> N \<Longrightarrow>
-              dist ((fm (Suc (diagseq i)) (y (Suc (diagseq i))))\<^sub>F t) (z t) < e" by auto
-            show "\<exists>N. \<forall>na\<ge>N. dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e "
-            proof (rule exI[where x="max N n"], safe)
-              fix na assume "max N n \<le> na"
-              hence  "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) =
-                      dist ((fm (Suc (diagseq na)) (y (Suc (diagseq na))))\<^sub>F t) (z t)" using t
-                by (subst index_shift[OF I]) auto
-              also have "\<dots> < e" using `max N n \<le> na` by (intro N) simp
-              finally show "dist ((fm n (y (Suc (diagseq na))))\<^sub>F t) (z t) < e" .
-            qed
-          qed
-          hence "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) ----> (finmap_of (Utn ` J n) z)\<^sub>F t"
-            by (simp add: tendsto_intros)
-        } ultimately
-        have "(\<lambda>i. fm n (y (Suc (diagseq i)))) ----> finmap_of (Utn ` J n) z"
-          by (rule tendsto_finmap)
-        hence "((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) ----> finmap_of (Utn ` J n) z"
-          by (intro lim_subseq) (simp add: subseq_def)
-        moreover
-        have "(\<forall>i. ((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) i \<in> K' n)"
-          apply (auto simp add: o_def intro!: fm_in_K' `1 \<le> n` le_SucI)
-          apply (rule le_trans)
-          apply (rule le_add2)
-          using seq_suble[OF subseq_diagseq]
-          apply auto
-          done
-        moreover
-        from `compact (K' n)` have "closed (K' n)" by (rule compact_imp_closed)
-        ultimately
-        have "finmap_of (Utn ` J n) z \<in> K' n"
-          unfolding closed_sequential_limits by blast
-        also have "finmap_of (Utn ` J n) z  = fm n (\<lambda>i. z (Utn i))"
-          unfolding finmap_eq_iff
-        proof clarsimp
-          fix i assume i: "i \<in> J n"
-          hence "from_nat_into (\<Union>n. J n) (Utn i) = i"
-            unfolding Utn_def
-            by (subst from_nat_into_to_nat_on[OF countable_UN_J]) auto
-          with i show "z (Utn i) = (fm n (\<lambda>i. z (Utn i)))\<^sub>F (Utn i)"
-            by (simp add: finmap_eq_iff fm_def compose_def)
-        qed
-        finally have "fm n (\<lambda>i. z (Utn i)) \<in> K' n" .
-        moreover
-        let ?J = "\<Union>n. J n"
-        have "(?J \<inter> J n) = J n" by auto
-        ultimately have "restrict (\<lambda>i. z (Utn i)) (?J \<inter> J n) \<in> K n"
-          unfolding K_def by (auto simp: proj_space space_PiM)
-        hence "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z' n" unfolding Z'_def
-          using J by (auto simp: prod_emb_def PiE_def extensional_def)
-        also have "\<dots> \<subseteq> Z n" using Z' by simp
-        finally have "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z n" .
-      } note in_Z = this
-      hence "(\<Inter>i\<in>{1..}. Z i) \<noteq> {}" by auto
-      hence "(\<Inter>i. Z i) \<noteq> {}" using Z INT_decseq_offset[OF `decseq Z`] by simp
-      thus False using Z by simp
+      hence "(\<lambda>i. (fm n (y (Suc (diagseq i))))\<^sub>F t) ----> (finmap_of (Utn ` J n) z)\<^sub>F t"
+        by (simp add: tendsto_intros)
+    } ultimately
+    have "(\<lambda>i. fm n (y (Suc (diagseq i)))) ----> finmap_of (Utn ` J n) z"
+      by (rule tendsto_finmap)
+    hence "((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) ----> finmap_of (Utn ` J n) z"
+      by (intro lim_subseq) (simp add: subseq_def)
+    moreover
+    have "(\<forall>i. ((\<lambda>i. fm n (y (Suc (diagseq i)))) o (\<lambda>i. i + n)) i \<in> K' n)"
+      apply (auto simp add: o_def intro!: fm_in_K' `1 \<le> n` le_SucI)
+      apply (rule le_trans)
+      apply (rule le_add2)
+      using seq_suble[OF subseq_diagseq]
+      apply auto
+      done
+    moreover
+    from `compact (K' n)` have "closed (K' n)" by (rule compact_imp_closed)
+    ultimately
+    have "finmap_of (Utn ` J n) z \<in> K' n"
+      unfolding closed_sequential_limits by blast
+    also have "finmap_of (Utn ` J n) z  = fm n (\<lambda>i. z (Utn i))"
+      unfolding finmap_eq_iff
+    proof clarsimp
+      fix i assume i: "i \<in> J n"
+      hence "from_nat_into (\<Union>n. J n) (Utn i) = i"
+        unfolding Utn_def
+        by (subst from_nat_into_to_nat_on[OF countable_UN_J]) auto
+      with i show "z (Utn i) = (fm n (\<lambda>i. z (Utn i)))\<^sub>F (Utn i)"
+        by (simp add: finmap_eq_iff fm_def compose_def)
     qed
-    ultimately show "(\<lambda>i. \<mu>G (Z i)) ----> 0"
-      using LIMSEQ_INF[of "\<lambda>i. \<mu>G (Z i)"] by simp
-  qed
-  then guess \<mu> .. note \<mu> = this
-  def f \<equiv> "finmap_of J B"
-  show "emeasure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = emeasure (lim\<^sub>B J P) (Pi\<^sub>E J B)"
-  proof (subst emeasure_extend_measure_Pair[OF limP_def, of I "\<lambda>_. borel" \<mu>])
-    show "positive (sets (lim\<^sub>B I P)) \<mu>" "countably_additive (sets (lim\<^sub>B I P)) \<mu>"
-      using \<mu> unfolding sets_limP sets_PiM_generator by (auto simp: measure_space_def)
-  next
-    show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> B \<in> J \<rightarrow> sets borel"
-      using assms by (auto simp: f_def)
-  next
-    fix J and X::"'i \<Rightarrow> 'a set"
-    show "prod_emb I (\<lambda>_. borel) J (Pi\<^sub>E J X) \<in> Pow (I \<rightarrow>\<^sub>E space borel)"
-      by (auto simp: prod_emb_def)
-    assume JX: "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> J \<rightarrow> sets borel"
-    hence "emb I J (Pi\<^sub>E J X) \<in> generator" using assms
-      by (intro generatorI[where J=J and X="Pi\<^sub>E J X"]) (auto intro: sets_PiM_I_finite)
-    hence "\<mu> (emb I J (Pi\<^sub>E J X)) = \<mu>G (emb I J (Pi\<^sub>E J X))" using \<mu> by simp
-    also have "\<dots> = emeasure (P J) (Pi\<^sub>E J X)"
-      using JX assms proj_sets
-      by (subst mu_G_eq) (auto simp: mu_G_eq limP_finite intro: sets_PiM_I_finite)
-    finally show "\<mu> (emb I J (Pi\<^sub>E J X)) = emeasure (P J) (Pi\<^sub>E J X)" .
-  next
-    show "emeasure (P J) (Pi\<^sub>E J B) = emeasure (limP J (\<lambda>_. borel) P) (Pi\<^sub>E J B)"
-      using assms by (simp add: f_def limP_finite Pi_def)
-  qed
-qed
+    finally have "fm n (\<lambda>i. z (Utn i)) \<in> K' n" .
+    moreover
+    let ?J = "\<Union>n. J n"
+    have "(?J \<inter> J n) = J n" by auto
+    ultimately have "restrict (\<lambda>i. z (Utn i)) (?J \<inter> J n) \<in> K n"
+      unfolding K_def by (auto simp: space_P space_PiM)
+    hence "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z' n" unfolding Z'_def
+      using J by (auto simp: prod_emb_def PiE_def extensional_def)
+    also have "\<dots> \<subseteq> Z n" using Z' by simp
+    finally have "restrict (\<lambda>i. z (Utn i)) ?J \<in> Z n" .
+  } note in_Z = this
+  hence "(\<Inter>i\<in>{1..}. Z i) \<noteq> {}" by auto
+  thus "(\<Inter>i. Z i) \<noteq> {}"
+    using INT_decseq_offset[OF antimonoI[OF Z_mono]] by simp
+qed fact+
+
+lemma measure_lim_emb:
+  "J \<subseteq> I \<Longrightarrow> finite J \<Longrightarrow> X \<in> sets (\<Pi>\<^sub>M i\<in>J. borel) \<Longrightarrow> measure lim (emb I J X) = measure (P J) X"
+  unfolding measure_def by (subst emeasure_lim_emb) auto
 
 end
 
@@ -548,70 +469,24 @@
 hide_const (open) domain
 hide_const (open) basis_finmap
 
-sublocale polish_projective \<subseteq> P!: prob_space "(lim\<^sub>B I P)"
+sublocale polish_projective \<subseteq> P!: prob_space lim
 proof
-  show "emeasure (lim\<^sub>B I P) (space (lim\<^sub>B I P)) = 1"
-  proof cases
-    assume "I = {}"
-    interpret prob_space "P {}" using proj_prob_space by simp
-    show ?thesis
-      by (simp add: space_PiM_empty limP_finite emeasure_space_1 `I = {}`)
-  next
-    assume "I \<noteq> {}"
-    then obtain i where "i \<in> I" by auto
-    interpret prob_space "P {i}" using proj_prob_space by simp
-    have R: "(space (lim\<^sub>B I P)) = (emb I {i} (Pi\<^sub>E {i} (\<lambda>_. space borel)))"
-      by (auto simp: prod_emb_def space_PiM)
-    moreover have "extensional {i} = space (P {i})" by (simp add: proj_space space_PiM PiE_def)
-    ultimately show ?thesis using `i \<in> I`
-      apply (subst R)
-      apply (subst emeasure_limB_emb_not_empty)
-      apply (auto simp: limP_finite emeasure_space_1 PiE_def)
-      done
-  qed
+  have *: "emb I {} {\<lambda>x. undefined} = space (\<Pi>\<^sub>M i\<in>I. borel)"
+    by (auto simp: prod_emb_def space_PiM)
+  interpret prob_space "P {}" 
+    using prob_space_P by simp
+  show "emeasure lim (space lim) = 1"
+    using emeasure_lim_emb[of "{}" "{\<lambda>x. undefined}"] emeasure_space_1
+    by (simp add: * PiM_empty space_P)
 qed
 
-context polish_projective begin
-
-lemma emeasure_limB_emb:
-  assumes X: "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
-  shows "emeasure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = emeasure (P J) (Pi\<^sub>E J B)"
-proof cases
-  interpret prob_space "P {}" using proj_prob_space by simp
-  assume "J = {}"
-  moreover have "emb I {} {\<lambda>x. undefined} = space (lim\<^sub>B I P)"
-    by (auto simp: space_PiM prod_emb_def)
-  moreover have "{\<lambda>x. undefined} = space (lim\<^sub>B {} P)"
-    by (auto simp: space_PiM prod_emb_def simp del: limP_finite)
-  ultimately show ?thesis
-    by (simp add: P.emeasure_space_1 limP_finite emeasure_space_1 del: space_limP)
-next
-  assume "J \<noteq> {}" with X show ?thesis
-    by (subst emeasure_limB_emb_not_empty) (auto simp: limP_finite)
-qed
-
-lemma measure_limB_emb:
-  assumes "J \<subseteq> I" "finite J" "\<forall>i\<in>J. B i \<in> sets borel"
-  shows "measure (lim\<^sub>B I P) (emb I J (Pi\<^sub>E J B)) = measure (P J) (Pi\<^sub>E J B)"
-proof -
-  interpret prob_space "P J" using proj_prob_space assms by simp
-  show ?thesis
-    using emeasure_limB_emb[OF assms]
-    unfolding emeasure_eq_measure limP_finite[OF `finite J` `J \<subseteq> I`] P.emeasure_eq_measure
-    by simp
-qed
-
-end
-
 locale polish_product_prob_space =
   product_prob_space "\<lambda>_. borel::('a::polish_space) measure" I for I::"'i set"
 
 sublocale polish_product_prob_space \<subseteq> P: polish_projective I "\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)"
 proof qed
 
-lemma (in polish_product_prob_space) limP_eq_PiM:
-  "I \<noteq> {} \<Longrightarrow> lim\<^sub>P I (\<lambda>_. borel) (\<lambda>J. PiM J (\<lambda>_. borel::('a) measure)) =
-    PiM I (\<lambda>_. borel)"
-  by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_limB_emb)
+lemma (in polish_product_prob_space) limP_eq_PiM: "lim = PiM I (\<lambda>_. borel)"
+  by (rule PiM_eq) (auto simp: emeasure_PiM emeasure_lim_emb)
 
 end