src/HOL/Probability/Infinite_Product_Measure.thy
author hoelzl
Mon, 23 Apr 2012 12:14:35 +0200
changeset 48565 05663f75964c
parent 47776 6b1c0a80a57a
child 48633 d31085f07f60
permissions -rw-r--r--
reworked Probability theory
     1 (*  Title:      HOL/Probability/Infinite_Product_Measure.thy
     2     Author:     Johannes Hölzl, TU München
     3 *)
     4 
     5 header {*Infinite Product Measure*}
     6 
     7 theory Infinite_Product_Measure
     8   imports Probability_Measure Caratheodory
     9 begin
    10 
    11 lemma sigma_sets_mono: assumes "A \<subseteq> sigma_sets X B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
    12 proof
    13   fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
    14     by induct (insert `A \<subseteq> sigma_sets X B`, auto intro: sigma_sets.intros)
    15 qed
    16 
    17 lemma sigma_sets_mono': assumes "A \<subseteq> B" shows "sigma_sets X A \<subseteq> sigma_sets X B"
    18 proof
    19   fix x assume "x \<in> sigma_sets X A" then show "x \<in> sigma_sets X B"
    20     by induct (insert `A \<subseteq> B`, auto intro: sigma_sets.intros)
    21 qed
    22 
    23 lemma sigma_sets_superset_generator: "A \<subseteq> sigma_sets X A"
    24   by (auto intro: sigma_sets.Basic)
    25 
    26 lemma (in product_sigma_finite)
    27   assumes IJ: "I \<inter> J = {}" "finite I" "finite J" and A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)"
    28   shows emeasure_fold_integral:
    29     "emeasure (Pi\<^isub>M (I \<union> J) M) A = (\<integral>\<^isup>+x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M)) \<partial>Pi\<^isub>M I M)" (is ?I)
    30     and emeasure_fold_measurable:
    31     "(\<lambda>x. emeasure (Pi\<^isub>M J M) (merge I x J -` A \<inter> space (Pi\<^isub>M J M))) \<in> borel_measurable (Pi\<^isub>M I M)" (is ?B)
    32 proof -
    33   interpret I: finite_product_sigma_finite M I by default fact
    34   interpret J: finite_product_sigma_finite M J by default fact
    35   interpret IJ: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" ..
    36   have merge: "(\<lambda>(x, y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M)"
    37     by (intro measurable_sets[OF _ A] measurable_merge assms)
    38 
    39   show ?I
    40     apply (subst distr_merge[symmetric, OF IJ])
    41     apply (subst emeasure_distr[OF measurable_merge[OF IJ(1)] A])
    42     apply (subst IJ.emeasure_pair_measure_alt[OF merge])
    43     apply (auto intro!: positive_integral_cong arg_cong2[where f=emeasure] simp: space_pair_measure)
    44     done
    45 
    46   show ?B
    47     using IJ.measurable_emeasure_Pair1[OF merge]
    48     by (simp add: vimage_compose[symmetric] comp_def space_pair_measure cong: measurable_cong)
    49 qed
    50 
    51 lemma restrict_extensional_sub[intro]: "A \<subseteq> B \<Longrightarrow> restrict f A \<in> extensional B"
    52   unfolding restrict_def extensional_def by auto
    53 
    54 lemma restrict_restrict[simp]: "restrict (restrict f A) B = restrict f (A \<inter> B)"
    55   unfolding restrict_def by (simp add: fun_eq_iff)
    56 
    57 lemma split_merge: "P (merge I x J y i) \<longleftrightarrow> (i \<in> I \<longrightarrow> P (x i)) \<and> (i \<in> J - I \<longrightarrow> P (y i)) \<and> (i \<notin> I \<union> J \<longrightarrow> P undefined)"
    58   unfolding merge_def by auto
    59 
    60 lemma extensional_merge_sub: "I \<union> J \<subseteq> K \<Longrightarrow> merge I x J y \<in> extensional K"
    61   unfolding merge_def extensional_def by auto
    62 
    63 lemma injective_vimage_restrict:
    64   assumes J: "J \<subseteq> I"
    65   and sets: "A \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" "B \<subseteq> (\<Pi>\<^isub>E i\<in>J. S i)" and ne: "(\<Pi>\<^isub>E i\<in>I. S i) \<noteq> {}"
    66   and eq: "(\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i) = (\<lambda>x. restrict x J) -` B \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
    67   shows "A = B"
    68 proof  (intro set_eqI)
    69   fix x
    70   from ne obtain y where y: "\<And>i. i \<in> I \<Longrightarrow> y i \<in> S i" by auto
    71   have "J \<inter> (I - J) = {}" by auto
    72   show "x \<in> A \<longleftrightarrow> x \<in> B"
    73   proof cases
    74     assume x: "x \<in> (\<Pi>\<^isub>E i\<in>J. S i)"
    75     have "x \<in> A \<longleftrightarrow> merge J x (I - J) y \<in> (\<lambda>x. restrict x J) -` A \<inter> (\<Pi>\<^isub>E i\<in>I. S i)"
    76       using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub split: split_merge)
    77     then show "x \<in> A \<longleftrightarrow> x \<in> B"
    78       using y x `J \<subseteq> I` by (auto simp add: Pi_iff extensional_restrict extensional_merge_sub eq split: split_merge)
    79   next
    80     assume "x \<notin> (\<Pi>\<^isub>E i\<in>J. S i)" with sets show "x \<in> A \<longleftrightarrow> x \<in> B" by auto
    81   qed
    82 qed
    83 
    84 lemma prod_algebraI_finite:
    85   "finite I \<Longrightarrow> (\<forall>i\<in>I. E i \<in> sets (M i)) \<Longrightarrow> (Pi\<^isub>E I E) \<in> prod_algebra I M"
    86   using prod_algebraI[of I I E M] prod_emb_PiE_same_index[of I E M, OF sets_into_space] by simp
    87 
    88 lemma Int_stable_PiE: "Int_stable {Pi\<^isub>E J E | E. \<forall>i\<in>I. E i \<in> sets (M i)}"
    89 proof (safe intro!: Int_stableI)
    90   fix E F assume "\<forall>i\<in>I. E i \<in> sets (M i)" "\<forall>i\<in>I. F i \<in> sets (M i)"
    91   then show "\<exists>G. Pi\<^isub>E J E \<inter> Pi\<^isub>E J F = Pi\<^isub>E J G \<and> (\<forall>i\<in>I. G i \<in> sets (M i))"
    92     by (auto intro!: exI[of _ "\<lambda>i. E i \<inter> F i"])
    93 qed
    94 
    95 lemma prod_emb_trans[simp]:
    96   "J \<subseteq> K \<Longrightarrow> K \<subseteq> L \<Longrightarrow> prod_emb L M K (prod_emb K M J X) = prod_emb L M J X"
    97   by (auto simp add: Int_absorb1 prod_emb_def)
    98 
    99 lemma prod_emb_Pi:
   100   assumes "X \<in> (\<Pi> j\<in>J. sets (M j))" "J \<subseteq> K"
   101   shows "prod_emb K M J (Pi\<^isub>E J X) = (\<Pi>\<^isub>E i\<in>K. if i \<in> J then X i else space (M i))"
   102   using assms space_closed
   103   by (auto simp: prod_emb_def Pi_iff split: split_if_asm) blast+
   104 
   105 lemma prod_emb_id:
   106   "B \<subseteq> (\<Pi>\<^isub>E i\<in>L. space (M i)) \<Longrightarrow> prod_emb L M L B = B"
   107   by (auto simp: prod_emb_def Pi_iff subset_eq extensional_restrict)
   108 
   109 lemma measurable_prod_emb[intro, simp]:
   110   "J \<subseteq> L \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> prod_emb L M J X \<in> sets (Pi\<^isub>M L M)"
   111   unfolding prod_emb_def space_PiM[symmetric]
   112   by (auto intro!: measurable_sets measurable_restrict measurable_component_singleton)
   113 
   114 lemma measurable_restrict_subset: "J \<subseteq> L \<Longrightarrow> (\<lambda>f. restrict f J) \<in> measurable (Pi\<^isub>M L M) (Pi\<^isub>M J M)"
   115   by (intro measurable_restrict measurable_component_singleton) auto
   116 
   117 lemma (in product_prob_space) distr_restrict:
   118   assumes "J \<noteq> {}" "J \<subseteq> K" "finite K"
   119   shows "(\<Pi>\<^isub>M i\<in>J. M i) = distr (\<Pi>\<^isub>M i\<in>K. M i) (\<Pi>\<^isub>M i\<in>J. M i) (\<lambda>f. restrict f J)" (is "?P = ?D")
   120 proof (rule measure_eqI_generator_eq)
   121   have "finite J" using `J \<subseteq> K` `finite K` by (auto simp add: finite_subset)
   122   interpret J: finite_product_prob_space M J proof qed fact
   123   interpret K: finite_product_prob_space M K proof qed fact
   124 
   125   let ?J = "{Pi\<^isub>E J E | E. \<forall>i\<in>J. E i \<in> sets (M i)}"
   126   let ?F = "\<lambda>i. \<Pi>\<^isub>E k\<in>J. space (M k)"
   127   let ?\<Omega> = "(\<Pi>\<^isub>E k\<in>J. space (M k))"
   128   show "Int_stable ?J"
   129     by (rule Int_stable_PiE)
   130   show "range ?F \<subseteq> ?J" "incseq ?F" "(\<Union>i. ?F i) = ?\<Omega>"
   131     using `finite J` by (auto intro!: prod_algebraI_finite)
   132   { fix i show "emeasure ?P (?F i) \<noteq> \<infinity>" by simp }
   133   show "?J \<subseteq> Pow ?\<Omega>" by (auto simp: Pi_iff dest: sets_into_space)
   134   show "sets (\<Pi>\<^isub>M i\<in>J. M i) = sigma_sets ?\<Omega> ?J" "sets ?D = sigma_sets ?\<Omega> ?J"
   135     using `finite J` by (simp_all add: sets_PiM prod_algebra_eq_finite Pi_iff)
   136   
   137   fix X assume "X \<in> ?J"
   138   then obtain E where [simp]: "X = Pi\<^isub>E J E" and E: "\<forall>i\<in>J. E i \<in> sets (M i)" by auto
   139   with `finite J` have X: "X \<in> sets (Pi\<^isub>M J M)" by auto
   140 
   141   have "emeasure ?P X = (\<Prod> i\<in>J. emeasure (M i) (E i))"
   142     using E by (simp add: J.measure_times)
   143   also have "\<dots> = (\<Prod> i\<in>J. emeasure (M i) (if i \<in> J then E i else space (M i)))"
   144     by simp
   145   also have "\<dots> = (\<Prod> i\<in>K. emeasure (M i) (if i \<in> J then E i else space (M i)))"
   146     using `finite K` `J \<subseteq> K`
   147     by (intro setprod_mono_one_left) (auto simp: M.emeasure_space_1)
   148   also have "\<dots> = emeasure (Pi\<^isub>M K M) (\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i))"
   149     using E by (simp add: K.measure_times)
   150   also have "(\<Pi>\<^isub>E i\<in>K. if i \<in> J then E i else space (M i)) = (\<lambda>f. restrict f J) -` Pi\<^isub>E J E \<inter> (\<Pi>\<^isub>E i\<in>K. space (M i))"
   151     using `J \<subseteq> K` sets_into_space E by (force simp:  Pi_iff split: split_if_asm)
   152   finally show "emeasure (Pi\<^isub>M J M) X = emeasure ?D X"
   153     using X `J \<subseteq> K` apply (subst emeasure_distr)
   154     by (auto intro!: measurable_restrict_subset simp: space_PiM)
   155 qed
   156 
   157 abbreviation (in product_prob_space)
   158   "emb L K X \<equiv> prod_emb L M K X"
   159 
   160 lemma (in product_prob_space) emeasure_prod_emb[simp]:
   161   assumes L: "J \<noteq> {}" "J \<subseteq> L" "finite L" and X: "X \<in> sets (Pi\<^isub>M J M)"
   162   shows "emeasure (Pi\<^isub>M L M) (emb L J X) = emeasure (Pi\<^isub>M J M) X"
   163   by (subst distr_restrict[OF L])
   164      (simp add: prod_emb_def space_PiM emeasure_distr measurable_restrict_subset L X)
   165 
   166 lemma (in product_prob_space) prod_emb_injective:
   167   assumes "J \<noteq> {}" "J \<subseteq> L" "finite J" and sets: "X \<in> sets (Pi\<^isub>M J M)" "Y \<in> sets (Pi\<^isub>M J M)"
   168   assumes "prod_emb L M J X = prod_emb L M J Y"
   169   shows "X = Y"
   170 proof (rule injective_vimage_restrict)
   171   show "X \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))" "Y \<subseteq> (\<Pi>\<^isub>E i\<in>J. space (M i))"
   172     using sets[THEN sets_into_space] by (auto simp: space_PiM)
   173   have "\<forall>i\<in>L. \<exists>x. x \<in> space (M i)"
   174     using M.not_empty by auto
   175   from bchoice[OF this]
   176   show "(\<Pi>\<^isub>E i\<in>L. space (M i)) \<noteq> {}" by auto
   177   show "(\<lambda>x. restrict x J) -` X \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i)) = (\<lambda>x. restrict x J) -` Y \<inter> (\<Pi>\<^isub>E i\<in>L. space (M i))"
   178     using `prod_emb L M J X = prod_emb L M J Y` by (simp add: prod_emb_def)
   179 qed fact
   180 
   181 definition (in product_prob_space) generator :: "('i \<Rightarrow> 'a) set set" where
   182   "generator = (\<Union>J\<in>{J. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I}. emb I J ` sets (Pi\<^isub>M J M))"
   183 
   184 lemma (in product_prob_space) generatorI':
   185   "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> emb I J X \<in> generator"
   186   unfolding generator_def by auto
   187 
   188 lemma (in product_prob_space) algebra_generator:
   189   assumes "I \<noteq> {}" shows "algebra (\<Pi>\<^isub>E i\<in>I. space (M i)) generator" (is "algebra ?\<Omega> ?G")
   190 proof
   191   let ?G = generator
   192   show "?G \<subseteq> Pow ?\<Omega>"
   193     by (auto simp: generator_def prod_emb_def)
   194   from `I \<noteq> {}` obtain i where "i \<in> I" by auto
   195   then show "{} \<in> ?G"
   196     by (auto intro!: exI[of _ "{i}"] image_eqI[where x="\<lambda>i. {}"]
   197              simp: sigma_sets.Empty generator_def prod_emb_def)
   198   from `i \<in> I` show "?\<Omega> \<in> ?G"
   199     by (auto intro!: exI[of _ "{i}"] image_eqI[where x="Pi\<^isub>E {i} (\<lambda>i. space (M i))"]
   200              simp: generator_def prod_emb_def)
   201   fix A assume "A \<in> ?G"
   202   then obtain JA XA where XA: "JA \<noteq> {}" "finite JA" "JA \<subseteq> I" "XA \<in> sets (Pi\<^isub>M JA M)" and A: "A = emb I JA XA"
   203     by (auto simp: generator_def)
   204   fix B assume "B \<in> ?G"
   205   then obtain JB XB where XB: "JB \<noteq> {}" "finite JB" "JB \<subseteq> I" "XB \<in> sets (Pi\<^isub>M JB M)" and B: "B = emb I JB XB"
   206     by (auto simp: generator_def)
   207   let ?RA = "emb (JA \<union> JB) JA XA"
   208   let ?RB = "emb (JA \<union> JB) JB XB"
   209   have *: "A - B = emb I (JA \<union> JB) (?RA - ?RB)" "A \<union> B = emb I (JA \<union> JB) (?RA \<union> ?RB)"
   210     using XA A XB B by auto
   211   show "A - B \<in> ?G" "A \<union> B \<in> ?G"
   212     unfolding * using XA XB by (safe intro!: generatorI') auto
   213 qed
   214 
   215 lemma (in product_prob_space) sets_PiM_generator:
   216   assumes "I \<noteq> {}" shows "sets (PiM I M) = sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
   217 proof
   218   show "sets (Pi\<^isub>M I M) \<subseteq> sigma_sets (\<Pi>\<^isub>E i\<in>I. space (M i)) generator"
   219     unfolding sets_PiM
   220   proof (safe intro!: sigma_sets_subseteq)
   221     fix A assume "A \<in> prod_algebra I M" with `I \<noteq> {}` show "A \<in> generator"
   222       by (auto intro!: generatorI' elim!: prod_algebraE)
   223   qed
   224 qed (auto simp: generator_def space_PiM[symmetric] intro!: sigma_sets_subset)
   225 
   226 lemma (in product_prob_space) generatorI:
   227   "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> A = emb I J X \<Longrightarrow> A \<in> generator"
   228   unfolding generator_def by auto
   229 
   230 definition (in product_prob_space)
   231   "\<mu>G A =
   232     (THE x. \<forall>J. J \<noteq> {} \<longrightarrow> finite J \<longrightarrow> J \<subseteq> I \<longrightarrow> (\<forall>X\<in>sets (Pi\<^isub>M J M). A = emb I J X \<longrightarrow> x = emeasure (Pi\<^isub>M J M) X))"
   233 
   234 lemma (in product_prob_space) \<mu>G_spec:
   235   assumes J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
   236   shows "\<mu>G A = emeasure (Pi\<^isub>M J M) X"
   237   unfolding \<mu>G_def
   238 proof (intro the_equality allI impI ballI)
   239   fix K Y assume K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "A = emb I K Y" "Y \<in> sets (Pi\<^isub>M K M)"
   240   have "emeasure (Pi\<^isub>M K M) Y = emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) K Y)"
   241     using K J by simp
   242   also have "emb (K \<union> J) K Y = emb (K \<union> J) J X"
   243     using K J by (simp add: prod_emb_injective[of "K \<union> J" I])
   244   also have "emeasure (Pi\<^isub>M (K \<union> J) M) (emb (K \<union> J) J X) = emeasure (Pi\<^isub>M J M) X"
   245     using K J by simp
   246   finally show "emeasure (Pi\<^isub>M J M) X = emeasure (Pi\<^isub>M K M) Y" ..
   247 qed (insert J, force)
   248 
   249 lemma (in product_prob_space) \<mu>G_eq:
   250   "J \<noteq> {} \<Longrightarrow> finite J \<Longrightarrow> J \<subseteq> I \<Longrightarrow> X \<in> sets (Pi\<^isub>M J M) \<Longrightarrow> \<mu>G (emb I J X) = emeasure (Pi\<^isub>M J M) X"
   251   by (intro \<mu>G_spec) auto
   252 
   253 lemma (in product_prob_space) generator_Ex:
   254   assumes *: "A \<in> generator"
   255   shows "\<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A = emb I J X \<and> \<mu>G A = emeasure (Pi\<^isub>M J M) X"
   256 proof -
   257   from * obtain J X where J: "J \<noteq> {}" "finite J" "J \<subseteq> I" "A = emb I J X" "X \<in> sets (Pi\<^isub>M J M)"
   258     unfolding generator_def by auto
   259   with \<mu>G_spec[OF this] show ?thesis by auto
   260 qed
   261 
   262 lemma (in product_prob_space) generatorE:
   263   assumes A: "A \<in> generator"
   264   obtains J X where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A" "\<mu>G A = emeasure (Pi\<^isub>M J M) X"
   265 proof -
   266   from generator_Ex[OF A] obtain X J where "J \<noteq> {}" "finite J" "J \<subseteq> I" "X \<in> sets (Pi\<^isub>M J M)" "emb I J X = A"
   267     "\<mu>G A = emeasure (Pi\<^isub>M J M) X" by auto
   268   then show thesis by (intro that) auto
   269 qed
   270 
   271 lemma (in product_prob_space) merge_sets:
   272   assumes "finite J" "finite K" "J \<inter> K = {}" and A: "A \<in> sets (Pi\<^isub>M (J \<union> K) M)" and x: "x \<in> space (Pi\<^isub>M J M)"
   273   shows "merge J x K -` A \<inter> space (Pi\<^isub>M K M) \<in> sets (Pi\<^isub>M K M)"
   274 proof -
   275   from sets_Pair1[OF
   276     measurable_merge[THEN measurable_sets, OF `J \<inter> K = {}`], OF A, of x] x
   277   show ?thesis
   278       by (simp add: space_pair_measure comp_def vimage_compose[symmetric])
   279 qed
   280 
   281 lemma (in product_prob_space) merge_emb:
   282   assumes "K \<subseteq> I" "J \<subseteq> I" and y: "y \<in> space (Pi\<^isub>M J M)"
   283   shows "(merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) =
   284     emb I (K - J) (merge J y (K - J) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M))"
   285 proof -
   286   have [simp]: "\<And>x J K L. merge J y K (restrict x L) = merge J y (K \<inter> L) x"
   287     by (auto simp: restrict_def merge_def)
   288   have [simp]: "\<And>x J K L. restrict (merge J y K x) L = merge (J \<inter> L) y (K \<inter> L) x"
   289     by (auto simp: restrict_def merge_def)
   290   have [simp]: "(I - J) \<inter> K = K - J" using `K \<subseteq> I` `J \<subseteq> I` by auto
   291   have [simp]: "(K - J) \<inter> (K \<union> J) = K - J" by auto
   292   have [simp]: "(K - J) \<inter> K = K - J" by auto
   293   from y `K \<subseteq> I` `J \<subseteq> I` show ?thesis
   294     by (simp split: split_merge add: prod_emb_def Pi_iff extensional_merge_sub set_eq_iff space_PiM)
   295        auto
   296 qed
   297 
   298 lemma (in product_prob_space) positive_\<mu>G: 
   299   assumes "I \<noteq> {}"
   300   shows "positive generator \<mu>G"
   301 proof -
   302   interpret G!: algebra "\<Pi>\<^isub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
   303   show ?thesis
   304   proof (intro positive_def[THEN iffD2] conjI ballI)
   305     from generatorE[OF G.empty_sets] guess J X . note this[simp]
   306     interpret J: finite_product_sigma_finite M J by default fact
   307     have "X = {}"
   308       by (rule prod_emb_injective[of J I]) simp_all
   309     then show "\<mu>G {} = 0" by simp
   310   next
   311     fix A assume "A \<in> generator"
   312     from generatorE[OF this] guess J X . note this[simp]
   313     interpret J: finite_product_sigma_finite M J by default fact
   314     show "0 \<le> \<mu>G A" by (simp add: emeasure_nonneg)
   315   qed
   316 qed
   317 
   318 lemma (in product_prob_space) additive_\<mu>G: 
   319   assumes "I \<noteq> {}"
   320   shows "additive generator \<mu>G"
   321 proof -
   322   interpret G!: algebra "\<Pi>\<^isub>E i\<in>I. space (M i)" generator by (rule algebra_generator) fact
   323   show ?thesis
   324   proof (intro additive_def[THEN iffD2] ballI impI)
   325     fix A assume "A \<in> generator" with generatorE guess J X . note J = this
   326     fix B assume "B \<in> generator" with generatorE guess K Y . note K = this
   327     assume "A \<inter> B = {}"
   328     have JK: "J \<union> K \<noteq> {}" "J \<union> K \<subseteq> I" "finite (J \<union> K)"
   329       using J K by auto
   330     interpret JK: finite_product_sigma_finite M "J \<union> K" by default fact
   331     have JK_disj: "emb (J \<union> K) J X \<inter> emb (J \<union> K) K Y = {}"
   332       apply (rule prod_emb_injective[of "J \<union> K" I])
   333       apply (insert `A \<inter> B = {}` JK J K)
   334       apply (simp_all add: Int prod_emb_Int)
   335       done
   336     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)"
   337       using J K by simp_all
   338     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))"
   339       by simp
   340     also have "\<dots> = emeasure (Pi\<^isub>M (J \<union> K) M) (emb (J \<union> K) J X \<union> emb (J \<union> K) K Y)"
   341       using JK J(1, 4) K(1, 4) by (simp add: \<mu>G_eq Un del: prod_emb_Un)
   342     also have "\<dots> = \<mu>G A + \<mu>G B"
   343       using J K JK_disj by (simp add: plus_emeasure[symmetric])
   344     finally show "\<mu>G (A \<union> B) = \<mu>G A + \<mu>G B" .
   345   qed
   346 qed
   347 
   348 lemma (in product_prob_space) emeasure_PiM_emb_not_empty:
   349   assumes X: "J \<noteq> {}" "J \<subseteq> I" "finite J" "\<forall>i\<in>J. X i \<in> sets (M i)"
   350   shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)"
   351 proof cases
   352   assume "finite I" with X show ?thesis by simp
   353 next
   354   let ?\<Omega> = "\<Pi>\<^isub>E i\<in>I. space (M i)"
   355   let ?G = generator
   356   assume "\<not> finite I"
   357   then have I_not_empty: "I \<noteq> {}" by auto
   358   interpret G!: algebra ?\<Omega> generator by (rule algebra_generator) fact
   359   note \<mu>G_mono =
   360     G.additive_increasing[OF positive_\<mu>G[OF I_not_empty] additive_\<mu>G[OF I_not_empty], THEN increasingD]
   361 
   362   { fix Z J assume J: "J \<noteq> {}" "finite J" "J \<subseteq> I" and Z: "Z \<in> ?G"
   363 
   364     from `infinite I` `finite J` obtain k where k: "k \<in> I" "k \<notin> J"
   365       by (metis rev_finite_subset subsetI)
   366     moreover from Z guess K' X' by (rule generatorE)
   367     moreover def K \<equiv> "insert k K'"
   368     moreover def X \<equiv> "emb K K' X'"
   369     ultimately have K: "K \<noteq> {}" "finite K" "K \<subseteq> I" "X \<in> sets (Pi\<^isub>M K M)" "Z = emb I K X"
   370       "K - J \<noteq> {}" "K - J \<subseteq> I" "\<mu>G Z = emeasure (Pi\<^isub>M K M) X"
   371       by (auto simp: subset_insertI)
   372 
   373     let ?M = "\<lambda>y. merge J y (K - J) -` emb (J \<union> K) K X \<inter> space (Pi\<^isub>M (K - J) M)"
   374     { fix y assume y: "y \<in> space (Pi\<^isub>M J M)"
   375       note * = merge_emb[OF `K \<subseteq> I` `J \<subseteq> I` y, of X]
   376       moreover
   377       have **: "?M y \<in> sets (Pi\<^isub>M (K - J) M)"
   378         using J K y by (intro merge_sets) auto
   379       ultimately
   380       have ***: "(merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<in> ?G"
   381         using J K by (intro generatorI) auto
   382       have "\<mu>G (merge J y (I - J) -` emb I K X \<inter> space (Pi\<^isub>M I M)) = emeasure (Pi\<^isub>M (K - J) M) (?M y)"
   383         unfolding * using K J by (subst \<mu>G_eq[OF _ _ _ **]) auto
   384       note * ** *** this }
   385     note merge_in_G = this
   386 
   387     have "finite (K - J)" using K by auto
   388 
   389     interpret J: finite_product_prob_space M J by default fact+
   390     interpret KmJ: finite_product_prob_space M "K - J" by default fact+
   391 
   392     have "\<mu>G Z = emeasure (Pi\<^isub>M (J \<union> (K - J)) M) (emb (J \<union> (K - J)) K X)"
   393       using K J by simp
   394     also have "\<dots> = (\<integral>\<^isup>+ x. emeasure (Pi\<^isub>M (K - J) M) (?M x) \<partial>Pi\<^isub>M J M)"
   395       using K J by (subst emeasure_fold_integral) auto
   396     also have "\<dots> = (\<integral>\<^isup>+ y. \<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M)) \<partial>Pi\<^isub>M J M)"
   397       (is "_ = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)")
   398     proof (intro positive_integral_cong)
   399       fix x assume x: "x \<in> space (Pi\<^isub>M J M)"
   400       with K merge_in_G(2)[OF this]
   401       show "emeasure (Pi\<^isub>M (K - J) M) (?M x) = \<mu>G (?MZ x)"
   402         unfolding `Z = emb I K X` merge_in_G(1)[OF x] by (subst \<mu>G_eq) auto
   403     qed
   404     finally have fold: "\<mu>G Z = (\<integral>\<^isup>+x. \<mu>G (?MZ x) \<partial>Pi\<^isub>M J M)" .
   405 
   406     { fix x assume x: "x \<in> space (Pi\<^isub>M J M)"
   407       then have "\<mu>G (?MZ x) \<le> 1"
   408         unfolding merge_in_G(4)[OF x] `Z = emb I K X`
   409         by (intro KmJ.measure_le_1 merge_in_G(2)[OF x]) }
   410     note le_1 = this
   411 
   412     let ?q = "\<lambda>y. \<mu>G (merge J y (I - J) -` Z \<inter> space (Pi\<^isub>M I M))"
   413     have "?q \<in> borel_measurable (Pi\<^isub>M J M)"
   414       unfolding `Z = emb I K X` using J K merge_in_G(3)
   415       by (simp add: merge_in_G  \<mu>G_eq emeasure_fold_measurable cong: measurable_cong)
   416     note this fold le_1 merge_in_G(3) }
   417   note fold = this
   418 
   419   have "\<exists>\<mu>. (\<forall>s\<in>?G. \<mu> s = \<mu>G s) \<and> measure_space ?\<Omega> (sigma_sets ?\<Omega> ?G) \<mu>"
   420   proof (rule G.caratheodory_empty_continuous[OF positive_\<mu>G additive_\<mu>G])
   421     fix A assume "A \<in> ?G"
   422     with generatorE guess J X . note JX = this
   423     interpret JK: finite_product_prob_space M J by default fact+
   424     from JX show "\<mu>G A \<noteq> \<infinity>" by simp
   425   next
   426     fix A assume A: "range A \<subseteq> ?G" "decseq A" "(\<Inter>i. A i) = {}"
   427     then have "decseq (\<lambda>i. \<mu>G (A i))"
   428       by (auto intro!: \<mu>G_mono simp: decseq_def)
   429     moreover
   430     have "(INF i. \<mu>G (A i)) = 0"
   431     proof (rule ccontr)
   432       assume "(INF i. \<mu>G (A i)) \<noteq> 0" (is "?a \<noteq> 0")
   433       moreover have "0 \<le> ?a"
   434         using A positive_\<mu>G[OF I_not_empty] by (auto intro!: INF_greatest simp: positive_def)
   435       ultimately have "0 < ?a" by auto
   436 
   437       have "\<forall>n. \<exists>J X. J \<noteq> {} \<and> finite J \<and> J \<subseteq> I \<and> X \<in> sets (Pi\<^isub>M J M) \<and> A n = emb I J X \<and> \<mu>G (A n) = emeasure (Pi\<^isub>M J M) X"
   438         using A by (intro allI generator_Ex) auto
   439       then obtain J' X' where J': "\<And>n. J' n \<noteq> {}" "\<And>n. finite (J' n)" "\<And>n. J' n \<subseteq> I" "\<And>n. X' n \<in> sets (Pi\<^isub>M (J' n) M)"
   440         and A': "\<And>n. A n = emb I (J' n) (X' n)"
   441         unfolding choice_iff by blast
   442       moreover def J \<equiv> "\<lambda>n. (\<Union>i\<le>n. J' i)"
   443       moreover def X \<equiv> "\<lambda>n. emb (J n) (J' n) (X' n)"
   444       ultimately have J: "\<And>n. J n \<noteq> {}" "\<And>n. finite (J n)" "\<And>n. J n \<subseteq> I" "\<And>n. X n \<in> sets (Pi\<^isub>M (J n) M)"
   445         by auto
   446       with A' have A_eq: "\<And>n. A n = emb I (J n) (X n)" "\<And>n. A n \<in> ?G"
   447         unfolding J_def X_def by (subst prod_emb_trans) (insert A, auto)
   448 
   449       have J_mono: "\<And>n m. n \<le> m \<Longrightarrow> J n \<subseteq> J m"
   450         unfolding J_def by force
   451 
   452       interpret J: finite_product_prob_space M "J i" for i by default fact+
   453 
   454       have a_le_1: "?a \<le> 1"
   455         using \<mu>G_spec[of "J 0" "A 0" "X 0"] J A_eq
   456         by (auto intro!: INF_lower2[of 0] J.measure_le_1)
   457 
   458       let ?M = "\<lambda>K Z y. merge K y (I - K) -` Z \<inter> space (Pi\<^isub>M I M)"
   459 
   460       { fix Z k assume Z: "range Z \<subseteq> ?G" "decseq Z" "\<forall>n. ?a / 2^k \<le> \<mu>G (Z n)"
   461         then have Z_sets: "\<And>n. Z n \<in> ?G" by auto
   462         fix J' assume J': "J' \<noteq> {}" "finite J'" "J' \<subseteq> I"
   463         interpret J': finite_product_prob_space M J' by default fact+
   464 
   465         let ?q = "\<lambda>n y. \<mu>G (?M J' (Z n) y)"
   466         let ?Q = "\<lambda>n. ?q n -` {?a / 2^(k+1) ..} \<inter> space (Pi\<^isub>M J' M)"
   467         { fix n
   468           have "?q n \<in> borel_measurable (Pi\<^isub>M J' M)"
   469             using Z J' by (intro fold(1)) auto
   470           then have "?Q n \<in> sets (Pi\<^isub>M J' M)"
   471             by (rule measurable_sets) auto }
   472         note Q_sets = this
   473 
   474         have "?a / 2^(k+1) \<le> (INF n. emeasure (Pi\<^isub>M J' M) (?Q n))"
   475         proof (intro INF_greatest)
   476           fix n
   477           have "?a / 2^k \<le> \<mu>G (Z n)" using Z by auto
   478           also have "\<dots> \<le> (\<integral>\<^isup>+ x. indicator (?Q n) x + ?a / 2^(k+1) \<partial>Pi\<^isub>M J' M)"
   479             unfolding fold(2)[OF J' `Z n \<in> ?G`]
   480           proof (intro positive_integral_mono)
   481             fix x assume x: "x \<in> space (Pi\<^isub>M J' M)"
   482             then have "?q n x \<le> 1 + 0"
   483               using J' Z fold(3) Z_sets by auto
   484             also have "\<dots> \<le> 1 + ?a / 2^(k+1)"
   485               using `0 < ?a` by (intro add_mono) auto
   486             finally have "?q n x \<le> 1 + ?a / 2^(k+1)" .
   487             with x show "?q n x \<le> indicator (?Q n) x + ?a / 2^(k+1)"
   488               by (auto split: split_indicator simp del: power_Suc)
   489           qed
   490           also have "\<dots> = emeasure (Pi\<^isub>M J' M) (?Q n) + ?a / 2^(k+1)"
   491             using `0 \<le> ?a` Q_sets J'.emeasure_space_1
   492             by (subst positive_integral_add) auto
   493           finally show "?a / 2^(k+1) \<le> emeasure (Pi\<^isub>M J' M) (?Q n)" using `?a \<le> 1`
   494             by (cases rule: ereal2_cases[of ?a "emeasure (Pi\<^isub>M J' M) (?Q n)"])
   495                (auto simp: field_simps)
   496         qed
   497         also have "\<dots> = emeasure (Pi\<^isub>M J' M) (\<Inter>n. ?Q n)"
   498         proof (intro INF_emeasure_decseq)
   499           show "range ?Q \<subseteq> sets (Pi\<^isub>M J' M)" using Q_sets by auto
   500           show "decseq ?Q"
   501             unfolding decseq_def
   502           proof (safe intro!: vimageI[OF refl])
   503             fix m n :: nat assume "m \<le> n"
   504             fix x assume x: "x \<in> space (Pi\<^isub>M J' M)"
   505             assume "?a / 2^(k+1) \<le> ?q n x"
   506             also have "?q n x \<le> ?q m x"
   507             proof (rule \<mu>G_mono)
   508               from fold(4)[OF J', OF Z_sets x]
   509               show "?M J' (Z n) x \<in> ?G" "?M J' (Z m) x \<in> ?G" by auto
   510               show "?M J' (Z n) x \<subseteq> ?M J' (Z m) x"
   511                 using `decseq Z`[THEN decseqD, OF `m \<le> n`] by auto
   512             qed
   513             finally show "?a / 2^(k+1) \<le> ?q m x" .
   514           qed
   515         qed simp
   516         finally have "(\<Inter>n. ?Q n) \<noteq> {}"
   517           using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
   518         then have "\<exists>w\<in>space (Pi\<^isub>M J' M). \<forall>n. ?a / 2 ^ (k + 1) \<le> ?q n w" by auto }
   519       note Ex_w = this
   520 
   521       let ?q = "\<lambda>k n y. \<mu>G (?M (J k) (A n) y)"
   522 
   523       have "\<forall>n. ?a / 2 ^ 0 \<le> \<mu>G (A n)" by (auto intro: INF_lower)
   524       from Ex_w[OF A(1,2) this J(1-3), of 0] guess w0 .. note w0 = this
   525 
   526       let ?P =
   527         "\<lambda>k wk w. w \<in> space (Pi\<^isub>M (J (Suc k)) M) \<and> restrict w (J k) = wk \<and>
   528           (\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n w)"
   529       def w \<equiv> "nat_rec w0 (\<lambda>k wk. Eps (?P k wk))"
   530 
   531       { fix k have w: "w k \<in> space (Pi\<^isub>M (J k) M) \<and>
   532           (\<forall>n. ?a / 2 ^ (k + 1) \<le> ?q k n (w k)) \<and> (k \<noteq> 0 \<longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1))"
   533         proof (induct k)
   534           case 0 with w0 show ?case
   535             unfolding w_def nat_rec_0 by auto
   536         next
   537           case (Suc k)
   538           then have wk: "w k \<in> space (Pi\<^isub>M (J k) M)" by auto
   539           have "\<exists>w'. ?P k (w k) w'"
   540           proof cases
   541             assume [simp]: "J k = J (Suc k)"
   542             show ?thesis
   543             proof (intro exI[of _ "w k"] conjI allI)
   544               fix n
   545               have "?a / 2 ^ (Suc k + 1) \<le> ?a / 2 ^ (k + 1)"
   546                 using `0 < ?a` `?a \<le> 1` by (cases ?a) (auto simp: field_simps)
   547               also have "\<dots> \<le> ?q k n (w k)" using Suc by auto
   548               finally show "?a / 2 ^ (Suc k + 1) \<le> ?q (Suc k) n (w k)" by simp
   549             next
   550               show "w k \<in> space (Pi\<^isub>M (J (Suc k)) M)"
   551                 using Suc by simp
   552               then show "restrict (w k) (J k) = w k"
   553                 by (simp add: extensional_restrict space_PiM)
   554             qed
   555           next
   556             assume "J k \<noteq> J (Suc k)"
   557             with J_mono[of k "Suc k"] have "J (Suc k) - J k \<noteq> {}" (is "?D \<noteq> {}") by auto
   558             have "range (\<lambda>n. ?M (J k) (A n) (w k)) \<subseteq> ?G"
   559               "decseq (\<lambda>n. ?M (J k) (A n) (w k))"
   560               "\<forall>n. ?a / 2 ^ (k + 1) \<le> \<mu>G (?M (J k) (A n) (w k))"
   561               using `decseq A` fold(4)[OF J(1-3) A_eq(2), of "w k" k] Suc
   562               by (auto simp: decseq_def)
   563             from Ex_w[OF this `?D \<noteq> {}`] J[of "Suc k"]
   564             obtain w' where w': "w' \<in> space (Pi\<^isub>M ?D M)"
   565               "\<forall>n. ?a / 2 ^ (Suc k + 1) \<le> \<mu>G (?M ?D (?M (J k) (A n) (w k)) w')" by auto
   566             let ?w = "merge (J k) (w k) ?D w'"
   567             have [simp]: "\<And>x. merge (J k) (w k) (I - J k) (merge ?D w' (I - ?D) x) =
   568               merge (J (Suc k)) ?w (I - (J (Suc k))) x"
   569               using J(3)[of "Suc k"] J(3)[of k] J_mono[of k "Suc k"]
   570               by (auto intro!: ext split: split_merge)
   571             have *: "\<And>n. ?M ?D (?M (J k) (A n) (w k)) w' = ?M (J (Suc k)) (A n) ?w"
   572               using w'(1) J(3)[of "Suc k"]
   573               by (auto simp: space_PiM split: split_merge intro!: extensional_merge_sub) force+
   574             show ?thesis
   575               apply (rule exI[of _ ?w])
   576               using w' J_mono[of k "Suc k"] wk unfolding *
   577               apply (auto split: split_merge intro!: extensional_merge_sub ext simp: space_PiM)
   578               apply (force simp: extensional_def)
   579               done
   580           qed
   581           then have "?P k (w k) (w (Suc k))"
   582             unfolding w_def nat_rec_Suc unfolding w_def[symmetric]
   583             by (rule someI_ex)
   584           then show ?case by auto
   585         qed
   586         moreover
   587         then have "w k \<in> space (Pi\<^isub>M (J k) M)" by auto
   588         moreover
   589         from w have "?a / 2 ^ (k + 1) \<le> ?q k k (w k)" by auto
   590         then have "?M (J k) (A k) (w k) \<noteq> {}"
   591           using positive_\<mu>G[OF I_not_empty, unfolded positive_def] `0 < ?a` `?a \<le> 1`
   592           by (cases ?a) (auto simp: divide_le_0_iff power_le_zero_eq)
   593         then obtain x where "x \<in> ?M (J k) (A k) (w k)" by auto
   594         then have "merge (J k) (w k) (I - J k) x \<in> A k" by auto
   595         then have "\<exists>x\<in>A k. restrict x (J k) = w k"
   596           using `w k \<in> space (Pi\<^isub>M (J k) M)`
   597           by (intro rev_bexI) (auto intro!: ext simp: extensional_def space_PiM)
   598         ultimately have "w k \<in> space (Pi\<^isub>M (J k) M)"
   599           "\<exists>x\<in>A k. restrict x (J k) = w k"
   600           "k \<noteq> 0 \<Longrightarrow> restrict (w k) (J (k - 1)) = w (k - 1)"
   601           by auto }
   602       note w = this
   603 
   604       { fix k l i assume "k \<le> l" "i \<in> J k"
   605         { fix l have "w k i = w (k + l) i"
   606           proof (induct l)
   607             case (Suc l)
   608             from `i \<in> J k` J_mono[of k "k + l"] have "i \<in> J (k + l)" by auto
   609             with w(3)[of "k + Suc l"]
   610             have "w (k + l) i = w (k + Suc l) i"
   611               by (auto simp: restrict_def fun_eq_iff split: split_if_asm)
   612             with Suc show ?case by simp
   613           qed simp }
   614         from this[of "l - k"] `k \<le> l` have "w l i = w k i" by simp }
   615       note w_mono = this
   616 
   617       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"
   618       { fix i k assume k: "i \<in> J k"
   619         have "w k i = w (LEAST k. i \<in> J k) i"
   620           by (intro w_mono Least_le k LeastI[of _ k])
   621         then have "w' i = w k i"
   622           unfolding w'_def using k by auto }
   623       note w'_eq = this
   624       have w'_simps1: "\<And>i. i \<notin> I \<Longrightarrow> w' i = undefined"
   625         using J by (auto simp: w'_def)
   626       have w'_simps2: "\<And>i. i \<notin> (\<Union>k. J k) \<Longrightarrow> i \<in> I \<Longrightarrow> w' i \<in> space (M i)"
   627         using J by (auto simp: w'_def intro!: someI_ex[OF M.not_empty[unfolded ex_in_conv[symmetric]]])
   628       { fix i assume "i \<in> I" then have "w' i \<in> space (M i)"
   629           using w(1) by (cases "i \<in> (\<Union>k. J k)") (force simp: w'_simps2 w'_eq space_PiM)+ }
   630       note w'_simps[simp] = w'_eq w'_simps1 w'_simps2 this
   631 
   632       have w': "w' \<in> space (Pi\<^isub>M I M)"
   633         using w(1) by (auto simp add: Pi_iff extensional_def space_PiM)
   634 
   635       { fix n
   636         have "restrict w' (J n) = w n" using w(1)
   637           by (auto simp add: fun_eq_iff restrict_def Pi_iff extensional_def space_PiM)
   638         with w[of n] obtain x where "x \<in> A n" "restrict x (J n) = restrict w' (J n)" by auto
   639         then have "w' \<in> A n" unfolding A_eq using w' by (auto simp: prod_emb_def space_PiM) }
   640       then have "w' \<in> (\<Inter>i. A i)" by auto
   641       with `(\<Inter>i. A i) = {}` show False by auto
   642     qed
   643     ultimately show "(\<lambda>i. \<mu>G (A i)) ----> 0"
   644       using LIMSEQ_ereal_INFI[of "\<lambda>i. \<mu>G (A i)"] by simp
   645   qed fact+
   646   then guess \<mu> .. note \<mu> = this
   647   show ?thesis
   648   proof (subst emeasure_extend_measure_Pair[OF PiM_def, of I M \<mu> J X])
   649     from assms show "(J \<noteq> {} \<or> I = {}) \<and> finite J \<and> J \<subseteq> I \<and> X \<in> (\<Pi> j\<in>J. sets (M j))"
   650       by (simp add: Pi_iff)
   651   next
   652     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))"
   653     then show "emb I J (Pi\<^isub>E J X) \<in> Pow (\<Pi>\<^isub>E i\<in>I. space (M i))"
   654       by (auto simp: Pi_iff prod_emb_def dest: sets_into_space)
   655     have "emb I J (Pi\<^isub>E J X) \<in> generator"
   656       using J `I \<noteq> {}` by (intro generatorI') auto
   657     then have "\<mu> (emb I J (Pi\<^isub>E J X)) = \<mu>G (emb I J (Pi\<^isub>E J X))"
   658       using \<mu> by simp
   659     also have "\<dots> = (\<Prod> j\<in>J. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
   660       using J  `I \<noteq> {}` by (subst \<mu>G_spec[OF _ _ _ refl]) (auto simp: emeasure_PiM Pi_iff)
   661     also have "\<dots> = (\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}.
   662       if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j)))"
   663       using J `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
   664     finally show "\<mu> (emb I J (Pi\<^isub>E J X)) = \<dots>" .
   665   next
   666     let ?F = "\<lambda>j. if j \<in> J then emeasure (M j) (X j) else emeasure (M j) (space (M j))"
   667     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)"
   668       using X `I \<noteq> {}` by (intro setprod_mono_one_right) (auto simp: M.emeasure_space_1)
   669     then show "(\<Prod>j\<in>J \<union> {i \<in> I. emeasure (M i) (space (M i)) \<noteq> 1}. ?F j) =
   670       emeasure (Pi\<^isub>M J M) (Pi\<^isub>E J X)"
   671       using X by (auto simp add: emeasure_PiM) 
   672   next
   673     show "positive (sets (Pi\<^isub>M I M)) \<mu>" "countably_additive (sets (Pi\<^isub>M I M)) \<mu>"
   674       using \<mu> unfolding sets_PiM_generator[OF `I \<noteq> {}`] by (auto simp: measure_space_def)
   675   qed
   676 qed
   677 
   678 sublocale product_prob_space \<subseteq> P: prob_space "Pi\<^isub>M I M"
   679 proof
   680   show "emeasure (Pi\<^isub>M I M) (space (Pi\<^isub>M I M)) = 1"
   681   proof cases
   682     assume "I = {}" then show ?thesis by (simp add: space_PiM_empty)
   683   next
   684     assume "I \<noteq> {}"
   685     then obtain i where "i \<in> I" by auto
   686     moreover then have "emb I {i} (\<Pi>\<^isub>E i\<in>{i}. space (M i)) = (space (Pi\<^isub>M I M))"
   687       by (auto simp: prod_emb_def space_PiM)
   688     ultimately show ?thesis
   689       using emeasure_PiM_emb_not_empty[of "{i}" "\<lambda>i. space (M i)"]
   690       by (simp add: emeasure_PiM emeasure_space_1)
   691   qed
   692 qed
   693 
   694 lemma (in product_prob_space) emeasure_PiM_emb:
   695   assumes X: "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
   696   shows "emeasure (Pi\<^isub>M I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. emeasure (M i) (X i))"
   697 proof cases
   698   assume "J = {}"
   699   moreover have "emb I {} {\<lambda>x. undefined} = space (Pi\<^isub>M I M)"
   700     by (auto simp: space_PiM prod_emb_def)
   701   ultimately show ?thesis
   702     by (simp add: space_PiM_empty P.emeasure_space_1)
   703 next
   704   assume "J \<noteq> {}" with X show ?thesis
   705     by (subst emeasure_PiM_emb_not_empty) (auto simp: emeasure_PiM)
   706 qed
   707 
   708 lemma (in product_prob_space) measure_PiM_emb:
   709   assumes "J \<subseteq> I" "finite J" "\<And>i. i \<in> J \<Longrightarrow> X i \<in> sets (M i)"
   710   shows "measure (PiM I M) (emb I J (Pi\<^isub>E J X)) = (\<Prod> i\<in>J. measure (M i) (X i))"
   711   using emeasure_PiM_emb[OF assms]
   712   unfolding emeasure_eq_measure M.emeasure_eq_measure by (simp add: setprod_ereal)
   713 
   714 lemma (in finite_product_prob_space) finite_measure_PiM_emb:
   715   "(\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)) \<Longrightarrow> measure (PiM I M) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))"
   716   using measure_PiM_emb[of I A] finite_index prod_emb_PiE_same_index[OF sets_into_space, of I A M]
   717   by auto
   718 
   719 subsection {* Sequence space *}
   720 
   721 locale sequence_space = product_prob_space M "UNIV :: nat set" for M
   722 
   723 lemma (in sequence_space) infprod_in_sets[intro]:
   724   fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets (M i)"
   725   shows "Pi UNIV E \<in> sets (Pi\<^isub>M UNIV M)"
   726 proof -
   727   have "Pi UNIV E = (\<Inter>i. emb UNIV {..i} (\<Pi>\<^isub>E j\<in>{..i}. E j))"
   728     using E E[THEN sets_into_space]
   729     by (auto simp: prod_emb_def Pi_iff extensional_def) blast
   730   with E show ?thesis by auto
   731 qed
   732 
   733 lemma (in sequence_space) measure_PiM_countable:
   734   fixes E :: "nat \<Rightarrow> 'a set" assumes E: "\<And>i. E i \<in> sets (M i)"
   735   shows "(\<lambda>n. \<Prod>i\<le>n. measure (M i) (E i)) ----> measure (Pi\<^isub>M UNIV M) (Pi UNIV E)"
   736 proof -
   737   let ?E = "\<lambda>n. emb UNIV {..n} (Pi\<^isub>E {.. n} E)"
   738   have "\<And>n. (\<Prod>i\<le>n. measure (M i) (E i)) = measure (Pi\<^isub>M UNIV M) (?E n)"
   739     using E by (simp add: measure_PiM_emb)
   740   moreover have "Pi UNIV E = (\<Inter>n. ?E n)"
   741     using E E[THEN sets_into_space]
   742     by (auto simp: prod_emb_def extensional_def Pi_iff) blast
   743   moreover have "range ?E \<subseteq> sets (Pi\<^isub>M UNIV M)"
   744     using E by auto
   745   moreover have "decseq ?E"
   746     by (auto simp: prod_emb_def Pi_iff decseq_def)
   747   ultimately show ?thesis
   748     by (simp add: finite_Lim_measure_decseq)
   749 qed
   750 
   751 end