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