| author | wenzelm | 
| Tue, 21 Sep 2010 21:53:15 +0200 | |
| changeset 39578 | b75164153c37 | 
| parent 39098 | 21e9bd6cf0a8 | 
| child 40859 | de0b30e6c2d2 | 
| permissions | -rw-r--r-- | 
| 35833 | 1 | theory Product_Measure | 
| 38656 | 2 | imports Lebesgue_Integration | 
| 35833 | 3 | begin | 
| 4 | ||
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 5 | definition "dynkin M \<longleftrightarrow> | 
| 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 6 | space M \<in> sets M \<and> | 
| 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 7 | (\<forall> A \<in> sets M. A \<subseteq> space M) \<and> | 
| 39098 | 8 | (\<forall> a \<in> sets M. \<forall> b \<in> sets M. a \<subseteq> b \<longrightarrow> b - a \<in> sets M) \<and> | 
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 9 | (\<forall>A. disjoint_family A \<and> range A \<subseteq> sets M \<longrightarrow> (\<Union>i::nat. A i) \<in> sets M)" | 
| 39080 | 10 | |
| 11 | lemma dynkinI: | |
| 12 | assumes "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M" | |
| 39094 | 13 | assumes "space M \<in> sets M" and "\<forall> b \<in> sets M. \<forall> a \<in> sets M. a \<subseteq> b \<longrightarrow> b - a \<in> sets M" | 
| 39080 | 14 |   assumes "\<And> a. (\<And> i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {})
 | 
| 15 | \<Longrightarrow> (\<And> i :: nat. a i \<in> sets M) \<Longrightarrow> UNION UNIV a \<in> sets M" | |
| 16 | shows "dynkin M" | |
| 39098 | 17 | using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff) | 
| 39080 | 18 | |
| 19 | lemma dynkin_subset: | |
| 20 | assumes "dynkin M" | |
| 21 | shows "\<And> A. A \<in> sets M \<Longrightarrow> A \<subseteq> space M" | |
| 22 | using assms unfolding dynkin_def by auto | |
| 23 | ||
| 24 | lemma dynkin_space: | |
| 25 | assumes "dynkin M" | |
| 26 | shows "space M \<in> sets M" | |
| 27 | using assms unfolding dynkin_def by auto | |
| 28 | ||
| 29 | lemma dynkin_diff: | |
| 30 | assumes "dynkin M" | |
| 39094 | 31 | shows "\<And> a b. \<lbrakk> a \<in> sets M ; b \<in> sets M ; a \<subseteq> b \<rbrakk> \<Longrightarrow> b - a \<in> sets M" | 
| 39080 | 32 | using assms unfolding dynkin_def by auto | 
| 33 | ||
| 39094 | 34 | lemma dynkin_empty: | 
| 35 | assumes "dynkin M" | |
| 36 |   shows "{} \<in> sets M"
 | |
| 37 | using dynkin_diff[OF assms dynkin_space[OF assms] dynkin_space[OF assms]] by auto | |
| 38 | ||
| 39080 | 39 | lemma dynkin_UN: | 
| 40 | assumes "dynkin M" | |
| 41 |   assumes "\<And> i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}"
 | |
| 39094 | 42 | assumes "\<And> i :: nat. a i \<in> sets M" | 
| 39080 | 43 | shows "UNION UNIV a \<in> sets M" | 
| 39098 | 44 | using assms by (auto simp: dynkin_def disjoint_family_on_def image_subset_iff) | 
| 39080 | 45 | |
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 46 | definition "Int_stable M \<longleftrightarrow> (\<forall> a \<in> sets M. \<forall> b \<in> sets M. a \<inter> b \<in> sets M)" | 
| 39080 | 47 | |
| 48 | lemma dynkin_trivial: | |
| 49 | shows "dynkin \<lparr> space = A, sets = Pow A \<rparr>" | |
| 50 | by (rule dynkinI) auto | |
| 51 | ||
| 39094 | 52 | lemma dynkin_lemma: | 
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 53 | fixes D :: "'a algebra" | 
| 39080 | 54 | assumes stab: "Int_stable E" | 
| 55 | and spac: "space E = space D" | |
| 56 | and subsED: "sets E \<subseteq> sets D" | |
| 57 | and subsDE: "sets D \<subseteq> sigma_sets (space E) (sets E)" | |
| 58 | and dyn: "dynkin D" | |
| 59 | shows "sigma (space E) (sets E) = D" | |
| 60 | proof - | |
| 61 |   def sets_\<delta>E == "\<Inter> {sets d | d :: 'a algebra. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}"
 | |
| 62 | def \<delta>E == "\<lparr> space = space E, sets = sets_\<delta>E \<rparr>" | |
| 63 |   have "\<lparr> space = space E, sets = Pow (space E) \<rparr> \<in> {d | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}"
 | |
| 64 | using dynkin_trivial spac subsED dynkin_subset[OF dyn] by fastsimp | |
| 65 |   hence not_empty: "{sets (d :: 'a algebra) | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d} \<noteq> {}"
 | |
| 66 | using exI[of "\<lambda> x. space x = space E \<and> dynkin x \<and> sets E \<subseteq> sets x" "\<lparr> space = space E, sets = Pow (space E) \<rparr>", simplified] | |
| 67 | by auto | |
| 39094 | 68 | have \<delta>E_D: "sets_\<delta>E \<subseteq> sets D" | 
| 39080 | 69 | unfolding sets_\<delta>E_def using assms by auto | 
| 70 | have \<delta>ynkin: "dynkin \<delta>E" | |
| 71 | proof (rule dynkinI, safe) | |
| 72 | fix A x assume asm: "A \<in> sets \<delta>E" "x \<in> A" | |
| 73 |     { fix d :: "('a, 'b) algebra_scheme" assume "A \<in> sets d" "dynkin d \<and> space d = space E"
 | |
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 74 | hence "A \<subseteq> space d" using dynkin_subset by auto } | 
| 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 75 | show "x \<in> space \<delta>E" using asm unfolding \<delta>E_def sets_\<delta>E_def using not_empty | 
| 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 76 | by simp (metis dynkin_subset in_mono mem_def) | 
| 39080 | 77 | next | 
| 78 | show "space \<delta>E \<in> sets \<delta>E" | |
| 79 | unfolding \<delta>E_def sets_\<delta>E_def | |
| 80 | using dynkin_space by fastsimp | |
| 81 | next | |
| 39094 | 82 | fix a b assume "a \<in> sets \<delta>E" "b \<in> sets \<delta>E" "a \<subseteq> b" | 
| 39080 | 83 | thus "b - a \<in> sets \<delta>E" | 
| 84 | unfolding \<delta>E_def sets_\<delta>E_def by (auto intro:dynkin_diff) | |
| 85 | next | |
| 86 |     fix a assume asm: "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" "\<And>i. a i \<in> sets \<delta>E"
 | |
| 87 | thus "UNION UNIV a \<in> sets \<delta>E" | |
| 88 | unfolding \<delta>E_def sets_\<delta>E_def apply (auto intro!:dynkin_UN[OF _ asm(1)]) | |
| 89 | by blast | |
| 90 | qed | |
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 91 | |
| 39080 | 92 |   def Dy == "\<lambda> d. {A | A. A \<in> sets_\<delta>E \<and> A \<inter> d \<in> sets_\<delta>E}"
 | 
| 93 |   { fix d assume dasm: "d \<in> sets_\<delta>E"
 | |
| 94 | have "dynkin \<lparr> space = space E, sets = Dy d \<rparr>" | |
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 95 | proof (rule dynkinI, safe, simp_all) | 
| 39080 | 96 | fix A x assume "A \<in> Dy d" "x \<in> A" | 
| 97 | thus "x \<in> space E" unfolding Dy_def sets_\<delta>E_def using not_empty | |
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 98 | by simp (metis dynkin_subset in_mono mem_def) | 
| 39080 | 99 | next | 
| 100 | show "space E \<in> Dy d" | |
| 101 | unfolding Dy_def \<delta>E_def sets_\<delta>E_def | |
| 102 | proof auto | |
| 103 | fix d assume asm: "dynkin d" "space d = space E" "sets E \<subseteq> sets d" | |
| 104 | hence "space d \<in> sets d" using dynkin_space[OF asm(1)] by auto | |
| 105 | thus "space E \<in> sets d" using asm by auto | |
| 106 | next | |
| 107 | fix da :: "'a algebra" assume asm: "dynkin da" "space da = space E" "sets E \<subseteq> sets da" | |
| 108 | have d: "d = space E \<inter> d" | |
| 109 | using dasm dynkin_subset[OF asm(1)] asm(2) dynkin_subset[OF \<delta>ynkin] | |
| 110 | unfolding \<delta>E_def by auto | |
| 111 | hence "space E \<inter> d \<in> sets \<delta>E" unfolding \<delta>E_def | |
| 112 | using dasm by auto | |
| 113 | have "sets \<delta>E \<subseteq> sets da" unfolding \<delta>E_def sets_\<delta>E_def using asm | |
| 114 | by auto | |
| 115 | thus "space E \<inter> d \<in> sets da" using dasm asm d dynkin_subset[OF \<delta>ynkin] | |
| 116 | unfolding \<delta>E_def by auto | |
| 117 | qed | |
| 118 | next | |
| 39094 | 119 | fix a b assume absm: "a \<in> Dy d" "b \<in> Dy d" "a \<subseteq> b" | 
| 39080 | 120 | hence "a \<in> sets \<delta>E" "b \<in> sets \<delta>E" | 
| 121 | unfolding Dy_def \<delta>E_def by auto | |
| 122 | hence *: "b - a \<in> sets \<delta>E" | |
| 39094 | 123 | using dynkin_diff[OF \<delta>ynkin] `a \<subseteq> b` by auto | 
| 39080 | 124 | have "a \<inter> d \<in> sets \<delta>E" "b \<inter> d \<in> sets \<delta>E" | 
| 125 | using absm unfolding Dy_def \<delta>E_def by auto | |
| 126 | hence "(b \<inter> d) - (a \<inter> d) \<in> sets \<delta>E" | |
| 39094 | 127 | using dynkin_diff[OF \<delta>ynkin] `a \<subseteq> b` by auto | 
| 39080 | 128 | hence **: "(b - a) \<inter> d \<in> sets \<delta>E" by (auto simp add:Diff_Int_distrib2) | 
| 129 | thus "b - a \<in> Dy d" | |
| 130 | using * ** unfolding Dy_def \<delta>E_def by auto | |
| 131 | next | |
| 132 |       fix a assume aasm: "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}" "\<And>i. a i \<in> Dy d"
 | |
| 39094 | 133 | hence "\<And> i. a i \<in> sets \<delta>E" | 
| 39080 | 134 | unfolding Dy_def \<delta>E_def by auto | 
| 135 | from dynkin_UN[OF \<delta>ynkin aasm(1) this] | |
| 136 | have *: "UNION UNIV a \<in> sets \<delta>E" by auto | |
| 137 | from aasm | |
| 138 | have aE: "\<forall> i. a i \<inter> d \<in> sets \<delta>E" | |
| 139 | unfolding Dy_def \<delta>E_def by auto | |
| 140 | from aasm | |
| 141 |       have "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> (a i \<inter> d) \<inter> (a j \<inter> d) = {}" by auto
 | |
| 142 | from dynkin_UN[OF \<delta>ynkin this] | |
| 143 | have "UNION UNIV (\<lambda> i. a i \<inter> d) \<in> sets \<delta>E" | |
| 144 | using aE by auto | |
| 145 | hence **: "UNION UNIV a \<inter> d \<in> sets \<delta>E" by auto | |
| 146 | from * ** show "UNION UNIV a \<in> Dy d" unfolding Dy_def \<delta>E_def by auto | |
| 147 | qed } note Dy_nkin = this | |
| 148 | have E_\<delta>E: "sets E \<subseteq> sets \<delta>E" | |
| 149 | unfolding \<delta>E_def sets_\<delta>E_def by auto | |
| 150 |   { fix d assume dasm: "d \<in> sets \<delta>E"
 | |
| 151 |     { fix e assume easm: "e \<in> sets E"
 | |
| 152 | hence deasm: "e \<in> sets \<delta>E" | |
| 153 | unfolding \<delta>E_def sets_\<delta>E_def by auto | |
| 154 | have subset: "Dy e \<subseteq> sets \<delta>E" | |
| 155 | unfolding Dy_def \<delta>E_def by auto | |
| 156 |       { fix e' assume e'asm: "e' \<in> sets E"
 | |
| 157 | have "e' \<inter> e \<in> sets E" | |
| 158 | using easm e'asm stab unfolding Int_stable_def by auto | |
| 159 | hence "e' \<inter> e \<in> sets \<delta>E" | |
| 160 | unfolding \<delta>E_def sets_\<delta>E_def by auto | |
| 161 | hence "e' \<in> Dy e" using e'asm unfolding Dy_def \<delta>E_def sets_\<delta>E_def by auto } | |
| 162 | hence E_Dy: "sets E \<subseteq> Dy e" by auto | |
| 163 |       have "\<lparr> space = space E, sets = Dy e \<rparr> \<in> {d | d. dynkin d \<and> space d = space E \<and> sets E \<subseteq> sets d}"
 | |
| 164 | using Dy_nkin[OF deasm[unfolded \<delta>E_def, simplified]] E_\<delta>E E_Dy by auto | |
| 165 | hence "sets_\<delta>E \<subseteq> Dy e" | |
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 166 | unfolding sets_\<delta>E_def by auto (metis E_Dy simps(1) simps(2) spac) | 
| 39080 | 167 | hence "sets \<delta>E = Dy e" using subset unfolding \<delta>E_def by auto | 
| 168 | hence "d \<inter> e \<in> sets \<delta>E" | |
| 169 | using dasm easm deasm unfolding Dy_def \<delta>E_def by auto | |
| 170 | hence "e \<in> Dy d" using deasm | |
| 171 | unfolding Dy_def \<delta>E_def | |
| 172 | by (auto simp add:Int_commute) } | |
| 173 | hence "sets E \<subseteq> Dy d" by auto | |
| 174 | hence "sets \<delta>E \<subseteq> Dy d" using Dy_nkin[OF dasm[unfolded \<delta>E_def, simplified]] | |
| 39082 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 175 | unfolding \<delta>E_def sets_\<delta>E_def | 
| 
54dbe0368dc6
changed definition of dynkin; replaces proofs by metis calles
 hoelzl parents: 
39080diff
changeset | 176 | by auto (metis `sets E <= Dy d` simps(1) simps(2) spac) | 
| 39080 | 177 | hence *: "sets \<delta>E = Dy d" | 
| 178 | unfolding Dy_def \<delta>E_def by auto | |
| 179 | fix a assume aasm: "a \<in> sets \<delta>E" | |
| 180 | hence "a \<inter> d \<in> sets \<delta>E" | |
| 181 | using * dasm unfolding Dy_def \<delta>E_def by auto } note \<delta>E_stab = this | |
| 39094 | 182 |   { fix A :: "nat \<Rightarrow> 'a set" assume Asm: "range A \<subseteq> sets \<delta>E" "\<And>A. A \<in> sets \<delta>E \<Longrightarrow> A \<subseteq> space \<delta>E"
 | 
| 183 | "\<And>a. a \<in> sets \<delta>E \<Longrightarrow> space \<delta>E - a \<in> sets \<delta>E" | |
| 184 |     "{} \<in> sets \<delta>E" "space \<delta>E \<in> sets \<delta>E"
 | |
| 185 |     let "?A i" = "A i \<inter> (\<Inter> j \<in> {..< i}. space \<delta>E - A j)"
 | |
| 186 |     { fix i :: nat
 | |
| 187 |       have *: "(\<Inter> j \<in> {..< i}. space \<delta>E - A j) \<inter> space \<delta>E \<in> sets \<delta>E"
 | |
| 188 | apply (induct i) | |
| 189 | using lessThan_Suc Asm \<delta>E_stab apply fastsimp | |
| 190 | apply (subst lessThan_Suc) | |
| 191 | apply (subst INT_insert) | |
| 192 | apply (subst Int_assoc) | |
| 193 | apply (subst \<delta>E_stab) | |
| 194 | using lessThan_Suc Asm \<delta>E_stab Asm | |
| 195 | apply (fastsimp simp add:Int_assoc dynkin_diff[OF \<delta>ynkin]) | |
| 196 | prefer 2 apply simp | |
| 197 | apply (rule dynkin_diff[OF \<delta>ynkin, of _ "space \<delta>E", OF _ dynkin_space[OF \<delta>ynkin]]) | |
| 198 | using Asm by auto | |
| 199 | have **: "\<And> i. A i \<subseteq> space \<delta>E" using Asm by auto | |
| 200 |       have "(\<Inter> j \<in> {..< i}. space \<delta>E - A j) \<subseteq> space \<delta>E \<or> (\<Inter> j \<in> {..< i}. A j) = UNIV \<and> i = 0"
 | |
| 201 | apply (cases i) | |
| 202 | using Asm ** dynkin_subset[OF \<delta>ynkin, of "A (i - 1)"] | |
| 203 | by auto | |
| 204 | hence Aisets: "?A i \<in> sets \<delta>E" | |
| 205 | apply (cases i) | |
| 206 | using Asm * apply fastsimp | |
| 207 | apply (rule \<delta>E_stab) | |
| 208 | using Asm * ** | |
| 209 | by (auto simp add:Int_absorb2) | |
| 210 | have "?A i = disjointed A i" unfolding disjointed_def | |
| 211 | atLeast0LessThan using Asm by auto | |
| 212 | hence "?A i = disjointed A i" "?A i \<in> sets \<delta>E" | |
| 213 | using Aisets by auto | |
| 214 | } note Ai = this | |
| 215 | from dynkin_UN[OF \<delta>ynkin _ this(2)] this disjoint_family_disjointed[of A] | |
| 216 | have "(\<Union> i. ?A i) \<in> sets \<delta>E" | |
| 217 | by (auto simp add:disjoint_family_on_def disjointed_def) | |
| 218 | hence "(\<Union> i. A i) \<in> sets \<delta>E" | |
| 219 | using Ai(1) UN_disjointed_eq[of A] by auto } note \<delta>E_UN = this | |
| 220 |   { fix a b assume asm: "a \<in> sets \<delta>E" "b \<in> sets \<delta>E"
 | |
| 221 |     let ?ab = "\<lambda> i. if (i::nat) = 0 then a else if i = 1 then b else {}"
 | |
| 222 | have *: "(\<Union> i. ?ab i) \<in> sets \<delta>E" | |
| 223 | apply (rule \<delta>E_UN) | |
| 224 | using asm \<delta>E_UN dynkin_empty[OF \<delta>ynkin] | |
| 225 | dynkin_subset[OF \<delta>ynkin] | |
| 226 | dynkin_space[OF \<delta>ynkin] | |
| 227 | dynkin_diff[OF \<delta>ynkin] by auto | |
| 228 | have "(\<Union> i. ?ab i) = a \<union> b" apply auto | |
| 229 | apply (case_tac "i = 0") | |
| 230 | apply auto | |
| 231 | apply (case_tac "i = 1") | |
| 232 | by auto | |
| 233 | hence "a \<union> b \<in> sets \<delta>E" using * by auto} note \<delta>E_Un = this | |
| 234 | have "sigma_algebra \<delta>E" | |
| 39080 | 235 | apply unfold_locales | 
| 39094 | 236 | using dynkin_subset[OF \<delta>ynkin] | 
| 237 | using dynkin_diff[OF \<delta>ynkin, of _ "space \<delta>E", OF _ dynkin_space[OF \<delta>ynkin]] | |
| 238 | using dynkin_diff[OF \<delta>ynkin, of "space \<delta>E" "space \<delta>E", OF dynkin_space[OF \<delta>ynkin] dynkin_space[OF \<delta>ynkin]] | |
| 239 | using dynkin_space[OF \<delta>ynkin] | |
| 240 | using \<delta>E_UN \<delta>E_Un | |
| 241 | by auto | |
| 242 | from sigma_algebra.sigma_subset[OF this E_\<delta>E] \<delta>E_D subsDE spac | |
| 243 | show ?thesis by (auto simp add:\<delta>E_def sigma_def) | |
| 244 | qed | |
| 245 | ||
| 246 | lemma measure_eq: | |
| 39096 | 247 | assumes fin: "\<mu> (space M) = \<nu> (space M)" "\<nu> (space M) < \<omega>" | 
| 39094 | 248 | assumes E: "M = sigma (space E) (sets E)" "Int_stable E" | 
| 249 | assumes eq: "\<And> e. e \<in> sets E \<Longrightarrow> \<mu> e = \<nu> e" | |
| 250 | assumes ms: "measure_space M \<mu>" "measure_space M \<nu>" | |
| 251 | assumes A: "A \<in> sets M" | |
| 252 | shows "\<mu> A = \<nu> A" | |
| 253 | proof - | |
| 254 | interpret M: measure_space M \<mu> | |
| 255 | using ms by simp | |
| 256 | interpret M': measure_space M \<nu> | |
| 257 | using ms by simp | |
| 258 | ||
| 259 |   let ?D_sets = "{A. A \<in> sets M \<and> \<mu> A = \<nu> A}"
 | |
| 260 | have \<delta>: "dynkin \<lparr> space = space M , sets = ?D_sets \<rparr>" | |
| 261 | proof (rule dynkinI, safe, simp_all) | |
| 262 | fix A x assume "A \<in> sets M \<and> \<mu> A = \<nu> A" "x \<in> A" | |
| 263 | thus "x \<in> space M" using assms M.sets_into_space by auto | |
| 264 | next | |
| 265 | show "\<mu> (space M) = \<nu> (space M)" | |
| 266 | using fin by auto | |
| 267 | next | |
| 268 | fix a b | |
| 269 | assume asm: "a \<in> sets M \<and> \<mu> a = \<nu> a" | |
| 270 | "b \<in> sets M \<and> \<mu> b = \<nu> b" "a \<subseteq> b" | |
| 271 | hence "a \<subseteq> space M" | |
| 272 | using M.sets_into_space by auto | |
| 273 | from M.measure_mono[OF this] | |
| 274 | have "\<mu> a \<le> \<mu> (space M)" | |
| 275 | using asm by auto | |
| 276 | hence afin: "\<mu> a < \<omega>" | |
| 277 | using fin by auto | |
| 278 | have *: "b = b - a \<union> a" using asm by auto | |
| 279 |     have **: "(b - a) \<inter> a = {}" using asm by auto
 | |
| 280 | have iv: "\<mu> (b - a) + \<mu> a = \<mu> b" | |
| 281 | using M.measure_additive[of "b - a" a] | |
| 282 | conjunct1[OF asm(1)] conjunct1[OF asm(2)] * ** | |
| 283 | by auto | |
| 284 | have v: "\<nu> (b - a) + \<nu> a = \<nu> b" | |
| 285 | using M'.measure_additive[of "b - a" a] | |
| 286 | conjunct1[OF asm(1)] conjunct1[OF asm(2)] * ** | |
| 287 | by auto | |
| 288 | from iv v have "\<mu> (b - a) = \<nu> (b - a)" using asm afin | |
| 289 | pinfreal_add_cancel_right[of "\<mu> (b - a)" "\<nu> a" "\<nu> (b - a)"] | |
| 290 | by auto | |
| 291 | thus "b - a \<in> sets M \<and> \<mu> (b - a) = \<nu> (b - a)" | |
| 292 | using asm by auto | |
| 293 | next | |
| 294 |     fix a assume "\<And>i j :: nat. i \<noteq> j \<Longrightarrow> a i \<inter> a j = {}"
 | |
| 295 | "\<And>i. a i \<in> sets M \<and> \<mu> (a i) = \<nu> (a i)" | |
| 296 | thus "(\<Union>x. a x) \<in> sets M \<and> \<mu> (\<Union>x. a x) = \<nu> (\<Union>x. a x)" | |
| 297 | using M.measure_countably_additive | |
| 298 | M'.measure_countably_additive | |
| 299 | M.countable_UN | |
| 300 | apply (auto simp add:disjoint_family_on_def image_def) | |
| 301 | apply (subst M.measure_countably_additive[symmetric]) | |
| 302 | apply (auto simp add:disjoint_family_on_def) | |
| 303 | apply (subst M'.measure_countably_additive[symmetric]) | |
| 304 | by (auto simp add:disjoint_family_on_def) | |
| 39080 | 305 | qed | 
| 39094 | 306 | have *: "sets E \<subseteq> ?D_sets" | 
| 307 | using eq E sigma_sets.Basic[of _ "sets E"] | |
| 308 | by (auto simp add:sigma_def) | |
| 309 | have **: "?D_sets \<subseteq> sets M" by auto | |
| 310 | have "M = \<lparr> space = space M , sets = ?D_sets \<rparr>" | |
| 311 | unfolding E(1) | |
| 312 | apply (rule dynkin_lemma[OF E(2)]) | |
| 313 | using eq E space_sigma \<delta> sigma_sets.Basic | |
| 314 | by (auto simp add:sigma_def) | |
| 315 | from subst[OF this, of "\<lambda> M. A \<in> sets M", OF A] | |
| 316 | show ?thesis by auto | |
| 317 | qed | |
| 39097 | 318 | (* | 
| 39094 | 319 | lemma | 
| 320 | assumes sfin: "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" "\<And> i :: nat. \<nu> (A i) < \<omega>" | |
| 321 | assumes A: "\<And> i. \<mu> (A i) = \<nu> (A i)" "\<And> i. A i \<subseteq> A (Suc i)" | |
| 322 | assumes E: "M = sigma (space E) (sets E)" "Int_stable E" | |
| 323 | assumes eq: "\<And> e. e \<in> sets E \<Longrightarrow> \<mu> e = \<nu> e" | |
| 324 | assumes ms: "measure_space (M :: 'a algebra) \<mu>" "measure_space M \<nu>" | |
| 325 | assumes B: "B \<in> sets M" | |
| 326 | shows "\<mu> B = \<nu> B" | |
| 327 | proof - | |
| 328 | interpret M: measure_space M \<mu> by (rule ms) | |
| 329 | interpret M': measure_space M \<nu> by (rule ms) | |
| 330 | have *: "M = \<lparr> space = space M, sets = sets M \<rparr>" by auto | |
| 331 |   { fix i :: nat
 | |
| 332 | have **: "M\<lparr> space := A i, sets := op \<inter> (A i) ` sets M \<rparr> = | |
| 333 | \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr>" | |
| 334 | by auto | |
| 335 | have mu_i: "measure_space \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr> \<mu>" | |
| 336 | using M.restricted_measure_space[of "A i", simplified **] | |
| 337 | sfin by auto | |
| 338 | have nu_i: "measure_space \<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr> \<nu>" | |
| 339 | using M'.restricted_measure_space[of "A i", simplified **] | |
| 340 | sfin by auto | |
| 341 | let ?M = "\<lparr> space = A i, sets = op \<inter> (A i) ` sets M \<rparr>" | |
| 342 | have "\<mu> (A i \<inter> B) = \<nu> (A i \<inter> B)" | |
| 343 | apply (rule measure_eq[of \<mu> ?M \<nu> "\<lparr> space = space E \<inter> A i, sets = op \<inter> (A i) ` sets E\<rparr>" "A i \<inter> B", simplified]) | |
| 344 | using assms nu_i mu_i | |
| 39095 | 345 | apply (auto simp add:image_def) (* TODO *) sorry | 
| 346 | show ?thesis sorry | |
| 39080 | 347 | qed | 
| 39097 | 348 | *) | 
| 38656 | 349 | definition prod_sets where | 
| 350 |   "prod_sets A B = {z. \<exists>x \<in> A. \<exists>y \<in> B. z = x \<times> y}"
 | |
| 351 | ||
| 35833 | 352 | definition | 
| 39088 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 353 | "prod_measure_space M1 M2 = sigma (space M1 \<times> space M2) (prod_sets (sets M1) (sets M2))" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 354 | |
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 355 | lemma | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 356 | fixes M1 :: "'a algebra" and M2 :: "'b algebra" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 357 | assumes "algebra M1" "algebra M2" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 358 | shows measureable_fst[intro!, simp]: | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 359 | "fst \<in> measurable (prod_measure_space M1 M2) M1" (is ?fst) | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 360 | and measureable_snd[intro!, simp]: | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 361 | "snd \<in> measurable (prod_measure_space M1 M2) M2" (is ?snd) | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 362 | proof - | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 363 | interpret M1: algebra M1 by fact | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 364 | interpret M2: algebra M2 by fact | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 365 | |
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 366 |   { fix X assume "X \<in> sets M1"
 | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 367 | then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. fst -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 368 | apply - apply (rule bexI[of _ X]) apply (rule bexI[of _ "space M2"]) | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 369 | using M1.sets_into_space by force+ } | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 370 | moreover | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 371 |   { fix X assume "X \<in> sets M2"
 | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 372 | then have "\<exists>X1\<in>sets M1. \<exists>X2\<in>sets M2. snd -` X \<inter> space M1 \<times> space M2 = X1 \<times> X2" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 373 | apply - apply (rule bexI[of _ "space M1"]) apply (rule bexI[of _ X]) | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 374 | using M2.sets_into_space by force+ } | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 375 | ultimately show ?fst ?snd | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 376 | by (force intro!: sigma_sets.Basic | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 377 | simp: measurable_def prod_measure_space_def prod_sets_def sets_sigma)+ | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 378 | qed | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 379 | |
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 380 | lemma (in sigma_algebra) measureable_prod: | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 381 | fixes M1 :: "'a algebra" and M2 :: "'b algebra" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 382 | assumes "algebra M1" "algebra M2" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 383 | shows "f \<in> measurable M (prod_measure_space M1 M2) \<longleftrightarrow> | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 384 | (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 385 | using assms proof (safe intro!: measurable_comp[where b="prod_measure_space M1 M2"]) | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 386 | interpret M1: algebra M1 by fact | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 387 | interpret M2: algebra M2 by fact | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 388 | assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 389 | |
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 390 | show "f \<in> measurable M (prod_measure_space M1 M2)" unfolding prod_measure_space_def | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 391 | proof (rule measurable_sigma) | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 392 | show "prod_sets (sets M1) (sets M2) \<subseteq> Pow (space M1 \<times> space M2)" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 393 | unfolding prod_sets_def using M1.sets_into_space M2.sets_into_space by auto | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 394 | show "f \<in> space M \<rightarrow> space M1 \<times> space M2" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 395 | using f s by (auto simp: mem_Times_iff measurable_def comp_def) | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 396 | fix A assume "A \<in> prod_sets (sets M1) (sets M2)" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 397 | then obtain B C where "B \<in> sets M1" "C \<in> sets M2" "A = B \<times> C" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 398 | unfolding prod_sets_def by auto | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 399 | moreover have "(fst \<circ> f) -` B \<inter> space M \<in> sets M" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 400 | using f `B \<in> sets M1` unfolding measurable_def by auto | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 401 | moreover have "(snd \<circ> f) -` C \<inter> space M \<in> sets M" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 402 | using s `C \<in> sets M2` unfolding measurable_def by auto | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 403 | moreover have "f -` A \<inter> space M = ((fst \<circ> f) -` B \<inter> space M) \<inter> ((snd \<circ> f) -` C \<inter> space M)" | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 404 | unfolding `A = B \<times> C` by (auto simp: vimage_Times) | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 405 | ultimately show "f -` A \<inter> space M \<in> sets M" by auto | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 406 | qed | 
| 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 407 | qed | 
| 35833 | 408 | |
| 409 | definition | |
| 39088 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 hoelzl parents: 
39082diff
changeset | 410 | "prod_measure M \<mu> N \<nu> = (\<lambda>A. measure_space.positive_integral M \<mu> (\<lambda>s0. \<nu> ((\<lambda>s1. (s0, s1)) -` A)))" | 
| 38656 | 411 | |
| 412 | lemma prod_setsI: "x \<in> A \<Longrightarrow> y \<in> B \<Longrightarrow> (x \<times> y) \<in> prod_sets A B" | |
| 413 | by (auto simp add: prod_sets_def) | |
| 35833 | 414 | |
| 38656 | 415 | lemma sigma_prod_sets_finite: | 
| 416 | assumes "finite A" and "finite B" | |
| 417 | shows "sigma_sets (A \<times> B) (prod_sets (Pow A) (Pow B)) = Pow (A \<times> B)" | |
| 418 | proof safe | |
| 419 | have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product) | |
| 35833 | 420 | |
| 38656 | 421 | fix x assume subset: "x \<subseteq> A \<times> B" | 
| 422 | hence "finite x" using fin by (rule finite_subset) | |
| 423 | from this subset show "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))" | |
| 424 | (is "x \<in> sigma_sets ?prod ?sets") | |
| 425 | proof (induct x) | |
| 426 | case empty show ?case by (rule sigma_sets.Empty) | |
| 427 | next | |
| 428 | case (insert a x) | |
| 429 |     hence "{a} \<in> sigma_sets ?prod ?sets" by (auto simp: prod_sets_def intro!: sigma_sets.Basic)
 | |
| 430 | moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto | |
| 431 | ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un) | |
| 432 | qed | |
| 433 | next | |
| 434 | fix x a b | |
| 435 | assume "x \<in> sigma_sets (A\<times>B) (prod_sets (Pow A) (Pow B))" and "(a, b) \<in> x" | |
| 436 | from sigma_sets_into_sp[OF _ this(1)] this(2) | |
| 437 | show "a \<in> A" and "b \<in> B" | |
| 438 | by (auto simp: prod_sets_def) | |
| 35833 | 439 | qed | 
| 440 | ||
| 38656 | 441 | lemma (in sigma_algebra) measurable_prod_sigma: | 
| 442 | assumes sa1: "sigma_algebra a1" and sa2: "sigma_algebra a2" | |
| 443 | assumes 1: "(fst o f) \<in> measurable M a1" and 2: "(snd o f) \<in> measurable M a2" | |
| 444 | shows "f \<in> measurable M (sigma ((space a1) \<times> (space a2)) | |
| 445 | (prod_sets (sets a1) (sets a2)))" | |
| 35977 | 446 | proof - | 
| 38656 | 447 | from 1 have fn1: "fst \<circ> f \<in> space M \<rightarrow> space a1" | 
| 448 | and q1: "\<forall>y\<in>sets a1. (fst \<circ> f) -` y \<inter> space M \<in> sets M" | |
| 449 | by (auto simp add: measurable_def) | |
| 450 | from 2 have fn2: "snd \<circ> f \<in> space M \<rightarrow> space a2" | |
| 451 | and q2: "\<forall>y\<in>sets a2. (snd \<circ> f) -` y \<inter> space M \<in> sets M" | |
| 452 | by (auto simp add: measurable_def) | |
| 453 | show ?thesis | |
| 454 | proof (rule measurable_sigma) | |
| 455 | show "prod_sets (sets a1) (sets a2) \<subseteq> Pow (space a1 \<times> space a2)" using sa1 sa2 | |
| 456 | by (auto simp add: prod_sets_def sigma_algebra_iff dest: algebra.space_closed) | |
| 457 | next | |
| 458 | show "f \<in> space M \<rightarrow> space a1 \<times> space a2" | |
| 459 | by (rule prod_final [OF fn1 fn2]) | |
| 460 | next | |
| 461 | fix z | |
| 462 | assume z: "z \<in> prod_sets (sets a1) (sets a2)" | |
| 463 | thus "f -` z \<inter> space M \<in> sets M" | |
| 464 | proof (auto simp add: prod_sets_def vimage_Times) | |
| 465 | fix x y | |
| 466 | assume x: "x \<in> sets a1" and y: "y \<in> sets a2" | |
| 467 | have "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M = | |
| 468 | ((fst \<circ> f) -` x \<inter> space M) \<inter> ((snd \<circ> f) -` y \<inter> space M)" | |
| 469 | by blast | |
| 470 | also have "... \<in> sets M" using x y q1 q2 | |
| 471 | by blast | |
| 472 | finally show "(fst \<circ> f) -` x \<inter> (snd \<circ> f) -` y \<inter> space M \<in> sets M" . | |
| 473 | qed | |
| 474 | qed | |
| 35977 | 475 | qed | 
| 35833 | 476 | |
| 38656 | 477 | lemma (in sigma_finite_measure) prod_measure_times: | 
| 478 | assumes "sigma_finite_measure N \<nu>" | |
| 479 | and "A1 \<in> sets M" "A2 \<in> sets N" | |
| 480 | shows "prod_measure M \<mu> N \<nu> (A1 \<times> A2) = \<mu> A1 * \<nu> A2" | |
| 481 | oops | |
| 35833 | 482 | |
| 38656 | 483 | lemma (in sigma_finite_measure) sigma_finite_prod_measure_space: | 
| 484 | assumes "sigma_finite_measure N \<nu>" | |
| 485 | shows "sigma_finite_measure (prod_measure_space M N) (prod_measure M \<mu> N \<nu>)" | |
| 486 | oops | |
| 487 | ||
| 488 | lemma (in finite_measure_space) finite_prod_measure_times: | |
| 489 | assumes "finite_measure_space N \<nu>" | |
| 490 | and "A1 \<in> sets M" "A2 \<in> sets N" | |
| 491 | shows "prod_measure M \<mu> N \<nu> (A1 \<times> A2) = \<mu> A1 * \<nu> A2" | |
| 492 | proof - | |
| 493 | interpret N: finite_measure_space N \<nu> by fact | |
| 494 |   have *: "\<And>x. \<nu> (Pair x -` (A1 \<times> A2)) * \<mu> {x} = (if x \<in> A1 then \<nu> A2 * \<mu> {x} else 0)"
 | |
| 495 | by (auto simp: vimage_Times comp_def) | |
| 496 | have "finite A1" | |
| 497 | using `A1 \<in> sets M` finite_space by (auto simp: sets_eq_Pow intro: finite_subset) | |
| 498 |   then have "\<mu> A1 = (\<Sum>x\<in>A1. \<mu> {x})" using `A1 \<in> sets M`
 | |
| 499 | by (auto intro!: measure_finite_singleton simp: sets_eq_Pow) | |
| 500 | then show ?thesis using `A1 \<in> sets M` | |
| 501 | unfolding prod_measure_def positive_integral_finite_eq_setsum * | |
| 502 | by (auto simp add: sets_eq_Pow setsum_right_distrib[symmetric] mult_commute setsum_cases[OF finite_space]) | |
| 35833 | 503 | qed | 
| 504 | ||
| 38656 | 505 | lemma (in finite_measure_space) finite_prod_measure_space: | 
| 506 | assumes "finite_measure_space N \<nu>" | |
| 507 | shows "prod_measure_space M N = \<lparr> space = space M \<times> space N, sets = Pow (space M \<times> space N) \<rparr>" | |
| 35977 | 508 | proof - | 
| 38656 | 509 | interpret N: finite_measure_space N \<nu> by fact | 
| 510 | show ?thesis using finite_space N.finite_space | |
| 511 | by (simp add: sigma_def prod_measure_space_def sigma_prod_sets_finite sets_eq_Pow N.sets_eq_Pow) | |
| 35833 | 512 | qed | 
| 513 | ||
| 38656 | 514 | lemma (in finite_measure_space) finite_measure_space_finite_prod_measure: | 
| 39097 | 515 |   fixes N :: "('c, 'd) algebra_scheme"
 | 
| 516 | assumes N: "finite_measure_space N \<nu>" | |
| 38656 | 517 | shows "finite_measure_space (prod_measure_space M N) (prod_measure M \<mu> N \<nu>)" | 
| 518 | unfolding finite_prod_measure_space[OF assms] | |
| 39097 | 519 | proof (rule finite_measure_spaceI, simp_all) | 
| 38656 | 520 | interpret N: finite_measure_space N \<nu> by fact | 
| 39097 | 521 | show "finite (space M \<times> space N)" using finite_space N.finite_space by auto | 
| 522 | show "prod_measure M \<mu> N \<nu> (space M \<times> space N) \<noteq> \<omega>" | |
| 523 | using finite_prod_measure_times[OF N top N.top] by simp | |
| 524 |   show "prod_measure M \<mu> N \<nu> {} = 0"
 | |
| 525 | using finite_prod_measure_times[OF N empty_sets N.empty_sets] by simp | |
| 38656 | 526 | |
| 39097 | 527 |   fix A B :: "('a * 'c) set" assume "A \<inter> B = {}" "A \<subseteq> space M \<times> space N" "B \<subseteq> space M \<times> space N"
 | 
| 528 | then show "prod_measure M \<mu> N \<nu> (A \<union> B) = prod_measure M \<mu> N \<nu> A + prod_measure M \<mu> N \<nu> B" | |
| 529 | apply (auto simp add: sets_eq_Pow prod_measure_def positive_integral_add[symmetric] | |
| 530 | intro!: positive_integral_cong) | |
| 531 | apply (subst N.measure_additive) | |
| 532 | apply (auto intro!: arg_cong[where f=\<mu>] simp: N.sets_eq_Pow sets_eq_Pow) | |
| 533 | done | |
| 38656 | 534 | qed | 
| 535 | ||
| 536 | lemma (in finite_measure_space) finite_measure_space_finite_prod_measure_alterantive: | |
| 537 | assumes N: "finite_measure_space N \<nu>" | |
| 538 | shows "finite_measure_space \<lparr> space = space M \<times> space N, sets = Pow (space M \<times> space N) \<rparr> (prod_measure M \<mu> N \<nu>)" | |
| 539 | (is "finite_measure_space ?M ?m") | |
| 540 | unfolding finite_prod_measure_space[OF N, symmetric] | |
| 541 | using finite_measure_space_finite_prod_measure[OF N] . | |
| 542 | ||
| 39097 | 543 | lemma prod_measure_times_finite: | 
| 544 | assumes fms: "finite_measure_space M \<mu>" "finite_measure_space N \<nu>" and a: "a \<in> space M \<times> space N" | |
| 545 |   shows "prod_measure M \<mu> N \<nu> {a} = \<mu> {fst a} * \<nu> {snd a}"
 | |
| 546 | proof (cases a) | |
| 547 | case (Pair b c) | |
| 548 |   hence a_eq: "{a} = {b} \<times> {c}" by simp
 | |
| 549 | ||
| 550 | interpret M: finite_measure_space M \<mu> by fact | |
| 551 | interpret N: finite_measure_space N \<nu> by fact | |
| 552 | ||
| 553 |   from finite_measure_space.finite_prod_measure_times[OF fms, of "{b}" "{c}"] M.sets_eq_Pow N.sets_eq_Pow a Pair
 | |
| 554 | show ?thesis unfolding a_eq by simp | |
| 555 | qed | |
| 556 | ||
| 39095 | 557 | end |