author | blanchet |
Mon, 21 May 2012 10:39:31 +0200 | |
changeset 47944 | e6b51fab96f7 |
parent 47694 | 05663f75964c |
child 49776 | 199d1d5bb17e |
permissions | -rw-r--r-- |
42146
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
1 |
(* Title: HOL/Probability/Binary_Product_Measure.thy |
42067 | 2 |
Author: Johannes Hölzl, TU München |
3 |
*) |
|
4 |
||
42146
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
5 |
header {*Binary product measures*} |
42067 | 6 |
|
42146
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
7 |
theory Binary_Product_Measure |
38656 | 8 |
imports Lebesgue_Integration |
35833 | 9 |
begin |
10 |
||
42950
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents:
42146
diff
changeset
|
11 |
lemma times_eq_iff: "A \<times> B = C \<times> D \<longleftrightarrow> A = C \<and> B = D \<or> ((A = {} \<or> B = {}) \<and> (C = {} \<or> D = {}))" |
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents:
42146
diff
changeset
|
12 |
by auto |
6e5c2a3c69da
move lemmas to Extended_Reals and Extended_Real_Limits
hoelzl
parents:
42146
diff
changeset
|
13 |
|
40859 | 14 |
lemma times_Int_times: "A \<times> B \<inter> C \<times> D = (A \<inter> C) \<times> (B \<inter> D)" |
15 |
by auto |
|
16 |
||
17 |
lemma Pair_vimage_times[simp]: "\<And>A B x. Pair x -` (A \<times> B) = (if x \<in> A then B else {})" |
|
18 |
by auto |
|
19 |
||
20 |
lemma rev_Pair_vimage_times[simp]: "\<And>A B y. (\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})" |
|
21 |
by auto |
|
22 |
||
23 |
lemma case_prod_distrib: "f (case x of (x, y) \<Rightarrow> g x y) = (case x of (x, y) \<Rightarrow> f (g x y))" |
|
24 |
by (cases x) simp |
|
25 |
||
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
26 |
lemma split_beta': "(\<lambda>(x,y). f x y) = (\<lambda>x. f (fst x) (snd x))" |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
27 |
by (auto simp: fun_eq_iff) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
28 |
|
40859 | 29 |
section "Binary products" |
30 |
||
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
31 |
definition pair_measure (infixr "\<Otimes>\<^isub>M" 80) where |
47694 | 32 |
"A \<Otimes>\<^isub>M B = measure_of (space A \<times> space B) |
33 |
{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} |
|
34 |
(\<lambda>X. \<integral>\<^isup>+x. (\<integral>\<^isup>+y. indicator X (x,y) \<partial>B) \<partial>A)" |
|
40859 | 35 |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
36 |
lemma space_pair_measure: |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
37 |
"space (A \<Otimes>\<^isub>M B) = space A \<times> space B" |
47694 | 38 |
unfolding pair_measure_def using space_closed[of A] space_closed[of B] |
39 |
by (intro space_measure_of) auto |
|
40 |
||
41 |
lemma sets_pair_measure: |
|
42 |
"sets (A \<Otimes>\<^isub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}" |
|
43 |
unfolding pair_measure_def using space_closed[of A] space_closed[of B] |
|
44 |
by (intro sets_measure_of) auto |
|
41095 | 45 |
|
47694 | 46 |
lemma pair_measureI[intro, simp]: |
47 |
"x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^isub>M B)" |
|
48 |
by (auto simp: sets_pair_measure) |
|
41095 | 49 |
|
47694 | 50 |
lemma measurable_pair_measureI: |
51 |
assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2" |
|
52 |
assumes 2: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> f -` (A \<times> B) \<inter> space M \<in> sets M" |
|
53 |
shows "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)" |
|
54 |
unfolding pair_measure_def using 1 2 |
|
55 |
by (intro measurable_measure_of) (auto dest: sets_into_space) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
56 |
|
47694 | 57 |
lemma measurable_fst[intro!, simp]: "fst \<in> measurable (M1 \<Otimes>\<^isub>M M2) M1" |
58 |
unfolding measurable_def |
|
59 |
proof safe |
|
60 |
fix A assume A: "A \<in> sets M1" |
|
61 |
from this[THEN sets_into_space] have "fst -` A \<inter> space M1 \<times> space M2 = A \<times> space M2" by auto |
|
62 |
with A show "fst -` A \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by (simp add: space_pair_measure) |
|
63 |
qed (simp add: space_pair_measure) |
|
40859 | 64 |
|
47694 | 65 |
lemma measurable_snd[intro!, simp]: "snd \<in> measurable (M1 \<Otimes>\<^isub>M M2) M2" |
66 |
unfolding measurable_def |
|
67 |
proof safe |
|
68 |
fix A assume A: "A \<in> sets M2" |
|
69 |
from this[THEN sets_into_space] have "snd -` A \<inter> space M1 \<times> space M2 = space M1 \<times> A" by auto |
|
70 |
with A show "snd -` A \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by (simp add: space_pair_measure) |
|
71 |
qed (simp add: space_pair_measure) |
|
72 |
||
73 |
lemma measurable_fst': "f \<in> measurable M (N \<Otimes>\<^isub>M P) \<Longrightarrow> (\<lambda>x. fst (f x)) \<in> measurable M N" |
|
74 |
using measurable_comp[OF _ measurable_fst] by (auto simp: comp_def) |
|
75 |
||
76 |
lemma measurable_snd': "f \<in> measurable M (N \<Otimes>\<^isub>M P) \<Longrightarrow> (\<lambda>x. snd (f x)) \<in> measurable M P" |
|
77 |
using measurable_comp[OF _ measurable_snd] by (auto simp: comp_def) |
|
40859 | 78 |
|
47694 | 79 |
lemma measurable_fst'': "f \<in> measurable M N \<Longrightarrow> (\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^isub>M P) N" |
80 |
using measurable_comp[OF measurable_fst _] by (auto simp: comp_def) |
|
81 |
||
82 |
lemma measurable_snd'': "f \<in> measurable M N \<Longrightarrow> (\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^isub>M M) N" |
|
83 |
using measurable_comp[OF measurable_snd _] by (auto simp: comp_def) |
|
84 |
||
85 |
lemma measurable_pair_iff: |
|
86 |
"f \<in> measurable M (M1 \<Otimes>\<^isub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2" |
|
87 |
proof safe |
|
88 |
assume f: "(fst \<circ> f) \<in> measurable M M1" and s: "(snd \<circ> f) \<in> measurable M M2" |
|
89 |
show "f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)" |
|
90 |
proof (rule measurable_pair_measureI) |
|
91 |
show "f \<in> space M \<rightarrow> space M1 \<times> space M2" |
|
92 |
using f s by (auto simp: mem_Times_iff measurable_def comp_def) |
|
93 |
fix A B assume "A \<in> sets M1" "B \<in> sets M2" |
|
94 |
moreover have "(fst \<circ> f) -` A \<inter> space M \<in> sets M" "(snd \<circ> f) -` B \<inter> space M \<in> sets M" |
|
95 |
using f `A \<in> sets M1` s `B \<in> sets M2` by (auto simp: measurable_sets) |
|
96 |
moreover have "f -` (A \<times> B) \<inter> space M = ((fst \<circ> f) -` A \<inter> space M) \<inter> ((snd \<circ> f) -` B \<inter> space M)" |
|
97 |
by (auto simp: vimage_Times) |
|
98 |
ultimately show "f -` (A \<times> B) \<inter> space M \<in> sets M" by auto |
|
40859 | 99 |
qed |
47694 | 100 |
qed auto |
40859 | 101 |
|
47694 | 102 |
lemma measurable_pair: |
103 |
"(fst \<circ> f) \<in> measurable M M1 \<Longrightarrow> (snd \<circ> f) \<in> measurable M M2 \<Longrightarrow> f \<in> measurable M (M1 \<Otimes>\<^isub>M M2)" |
|
104 |
unfolding measurable_pair_iff by simp |
|
40859 | 105 |
|
47694 | 106 |
lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) (M2 \<Otimes>\<^isub>M M1)" |
107 |
proof (rule measurable_pair_measureI) |
|
108 |
fix A B assume "A \<in> sets M2" "B \<in> sets M1" |
|
109 |
moreover then have "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space (M1 \<Otimes>\<^isub>M M2) = B \<times> A" |
|
110 |
by (auto dest: sets_into_space simp: space_pair_measure) |
|
111 |
ultimately show "(\<lambda>(x, y). (y, x)) -` (A \<times> B) \<inter> space (M1 \<Otimes>\<^isub>M M2) \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
|
112 |
by auto |
|
113 |
qed (auto simp add: space_pair_measure) |
|
114 |
||
115 |
lemma measurable_pair_swap: |
|
116 |
assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^isub>M M1) M" |
|
117 |
proof - |
|
118 |
have "(\<lambda>x. f (case x of (x, y) \<Rightarrow> (y, x))) = (\<lambda>(x, y). f (y, x))" by auto |
|
119 |
then show ?thesis |
|
120 |
using measurable_comp[OF measurable_pair_swap' f] by (simp add: comp_def) |
|
40859 | 121 |
qed |
122 |
||
47694 | 123 |
lemma measurable_pair_swap_iff: |
124 |
"f \<in> measurable (M2 \<Otimes>\<^isub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" |
|
125 |
using measurable_pair_swap[of "\<lambda>(x,y). f (y, x)"] |
|
126 |
by (auto intro!: measurable_pair_swap) |
|
40859 | 127 |
|
47694 | 128 |
lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^isub>M M2)" |
129 |
proof (rule measurable_pair_measureI) |
|
130 |
fix A B assume "A \<in> sets M1" "B \<in> sets M2" |
|
131 |
moreover then have "Pair x -` (A \<times> B) \<inter> space M2 = (if x \<in> A then B else {})" |
|
132 |
by (auto dest: sets_into_space simp: space_pair_measure) |
|
133 |
ultimately show "Pair x -` (A \<times> B) \<inter> space M2 \<in> sets M2" |
|
134 |
by auto |
|
135 |
qed (auto simp add: space_pair_measure) |
|
40859 | 136 |
|
47694 | 137 |
lemma sets_Pair1: assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "Pair x -` A \<in> sets M2" |
40859 | 138 |
proof - |
47694 | 139 |
have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})" |
140 |
using A[THEN sets_into_space] by (auto simp: space_pair_measure) |
|
141 |
also have "\<dots> \<in> sets M2" |
|
142 |
using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: split_if_asm) |
|
143 |
finally show ?thesis . |
|
40859 | 144 |
qed |
145 |
||
47694 | 146 |
lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^isub>M M2)" |
147 |
using measurable_comp[OF measurable_Pair1' measurable_pair_swap', of y M2 M1] |
|
148 |
by (simp add: comp_def) |
|
40859 | 149 |
|
47694 | 150 |
lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1" |
151 |
proof - |
|
152 |
have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})" |
|
153 |
using A[THEN sets_into_space] by (auto simp: space_pair_measure) |
|
154 |
also have "\<dots> \<in> sets M1" |
|
155 |
using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: split_if_asm) |
|
156 |
finally show ?thesis . |
|
40859 | 157 |
qed |
158 |
||
47694 | 159 |
lemma measurable_Pair2: |
160 |
assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" and x: "x \<in> space M1" |
|
161 |
shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M" |
|
162 |
using measurable_comp[OF measurable_Pair1' f, OF x] |
|
163 |
by (simp add: comp_def) |
|
164 |
||
165 |
lemma measurable_Pair1: |
|
166 |
assumes f: "f \<in> measurable (M1 \<Otimes>\<^isub>M M2) M" and y: "y \<in> space M2" |
|
40859 | 167 |
shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M" |
47694 | 168 |
using measurable_comp[OF measurable_Pair2' f, OF y] |
169 |
by (simp add: comp_def) |
|
40859 | 170 |
|
47694 | 171 |
lemma Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}" |
40859 | 172 |
unfolding Int_stable_def |
47694 | 173 |
by safe (auto simp add: times_Int_times) |
40859 | 174 |
|
175 |
lemma finite_measure_cut_measurable: |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
176 |
assumes "sigma_finite_measure M1" "finite_measure M2" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
177 |
assumes "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
47694 | 178 |
shows "(\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1" |
40859 | 179 |
(is "?s Q \<in> _") |
180 |
proof - |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
181 |
interpret M1: sigma_finite_measure M1 by fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
182 |
interpret M2: finite_measure M2 by fact |
47694 | 183 |
let ?\<Omega> = "space M1 \<times> space M2" and ?D = "{A\<in>sets (M1 \<Otimes>\<^isub>M M2). ?s A \<in> borel_measurable M1}" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
184 |
note space_pair_measure[simp] |
47694 | 185 |
interpret dynkin_system ?\<Omega> ?D |
40859 | 186 |
proof (intro dynkin_systemI) |
47694 | 187 |
fix A assume "A \<in> ?D" then show "A \<subseteq> ?\<Omega>" |
188 |
using sets_into_space[of A "M1 \<Otimes>\<^isub>M M2"] by simp |
|
40859 | 189 |
next |
47694 | 190 |
from top show "?\<Omega> \<in> ?D" |
191 |
by (auto simp add: if_distrib intro!: measurable_If) |
|
40859 | 192 |
next |
47694 | 193 |
fix A assume "A \<in> ?D" |
194 |
with sets_into_space have "\<And>x. emeasure M2 (Pair x -` (?\<Omega> - A)) = |
|
195 |
(if x \<in> space M1 then emeasure M2 (space M2) - ?s A x else 0)" |
|
196 |
by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1) |
|
197 |
with `A \<in> ?D` top show "?\<Omega> - A \<in> ?D" |
|
198 |
by (auto intro!: measurable_If) |
|
40859 | 199 |
next |
47694 | 200 |
fix F :: "nat \<Rightarrow> ('a\<times>'b) set" assume "disjoint_family F" "range F \<subseteq> ?D" |
201 |
moreover then have "\<And>x. emeasure M2 (\<Union>i. Pair x -` F i) = (\<Sum>i. ?s (F i) x)" |
|
202 |
by (intro suminf_emeasure[symmetric]) (auto simp: disjoint_family_on_def sets_Pair1) |
|
203 |
ultimately show "(\<Union>i. F i) \<in> ?D" |
|
204 |
by (auto simp: vimage_UN intro!: borel_measurable_psuminf) |
|
40859 | 205 |
qed |
47694 | 206 |
let ?G = "{a \<times> b | a b. a \<in> sets M1 \<and> b \<in> sets M2}" |
207 |
have "sigma_sets ?\<Omega> ?G = ?D" |
|
208 |
proof (rule dynkin_lemma) |
|
209 |
show "?G \<subseteq> ?D" |
|
210 |
by (auto simp: if_distrib Int_def[symmetric] intro!: measurable_If) |
|
211 |
qed (auto simp: sets_pair_measure Int_stable_pair_measure_generator) |
|
212 |
with `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` have "Q \<in> ?D" |
|
213 |
by (simp add: sets_pair_measure[symmetric]) |
|
40859 | 214 |
then show "?s Q \<in> borel_measurable M1" by simp |
215 |
qed |
|
216 |
||
47694 | 217 |
subsection {* Binary products of $\sigma$-finite emeasure spaces *} |
40859 | 218 |
|
47694 | 219 |
locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2 |
220 |
for M1 :: "'a measure" and M2 :: "'b measure" |
|
40859 | 221 |
|
47694 | 222 |
lemma sets_pair_measure_cong[cong]: |
223 |
"sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^isub>M M2) = sets (M1' \<Otimes>\<^isub>M M2')" |
|
224 |
unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq) |
|
225 |
||
226 |
lemma (in pair_sigma_finite) measurable_emeasure_Pair1: |
|
227 |
assumes Q: "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1" (is "?s Q \<in> _") |
|
40859 | 228 |
proof - |
47694 | 229 |
from M2.sigma_finite_disjoint guess F . note F = this |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
230 |
then have F_sets: "\<And>i. F i \<in> sets M2" by auto |
47694 | 231 |
have M1: "sigma_finite_measure M1" .. |
46731 | 232 |
let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q" |
40859 | 233 |
{ fix i |
234 |
have [simp]: "space M1 \<times> F i \<inter> space M1 \<times> space M2 = space M1 \<times> F i" |
|
47694 | 235 |
using F sets_into_space by auto |
236 |
let ?R = "density M2 (indicator (F i))" |
|
237 |
have "(\<lambda>x. emeasure ?R (Pair x -` (space M1 \<times> space ?R \<inter> Q))) \<in> borel_measurable M1" |
|
40859 | 238 |
proof (intro finite_measure_cut_measurable[OF M1]) |
47694 | 239 |
show "finite_measure ?R" |
240 |
using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq) |
|
241 |
show "(space M1 \<times> space ?R) \<inter> Q \<in> sets (M1 \<Otimes>\<^isub>M ?R)" |
|
242 |
using Q by (simp add: Int) |
|
40859 | 243 |
qed |
47694 | 244 |
moreover have "\<And>x. emeasure ?R (Pair x -` (space M1 \<times> space ?R \<inter> Q)) |
245 |
= emeasure M2 (F i \<inter> Pair x -` (space M1 \<times> space ?R \<inter> Q))" |
|
246 |
using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1) |
|
247 |
moreover have "\<And>x. F i \<inter> Pair x -` (space M1 \<times> space ?R \<inter> Q) = ?C x i" |
|
248 |
using sets_into_space[OF Q] by (auto simp: space_pair_measure) |
|
249 |
ultimately have "(\<lambda>x. emeasure M2 (?C x i)) \<in> borel_measurable M1" |
|
40859 | 250 |
by simp } |
251 |
moreover |
|
252 |
{ fix x |
|
47694 | 253 |
have "(\<Sum>i. emeasure M2 (?C x i)) = emeasure M2 (\<Union>i. ?C x i)" |
254 |
proof (intro suminf_emeasure) |
|
40859 | 255 |
show "range (?C x) \<subseteq> sets M2" |
47694 | 256 |
using F `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` by (auto intro!: sets_Pair1) |
40859 | 257 |
have "disjoint_family F" using F by auto |
258 |
show "disjoint_family (?C x)" |
|
259 |
by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto |
|
260 |
qed |
|
261 |
also have "(\<Union>i. ?C x i) = Pair x -` Q" |
|
47694 | 262 |
using F sets_into_space[OF `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)`] |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
263 |
by (auto simp: space_pair_measure) |
47694 | 264 |
finally have "emeasure M2 (Pair x -` Q) = (\<Sum>i. emeasure M2 (?C x i))" |
40859 | 265 |
by simp } |
47694 | 266 |
ultimately show ?thesis using `Q \<in> sets (M1 \<Otimes>\<^isub>M M2)` F_sets |
267 |
by auto |
|
40859 | 268 |
qed |
269 |
||
47694 | 270 |
lemma (in pair_sigma_finite) measurable_emeasure_Pair2: |
271 |
assumes Q: "Q \<in> sets (M1 \<Otimes>\<^isub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2" |
|
40859 | 272 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
273 |
interpret Q: pair_sigma_finite M2 M1 by default |
47694 | 274 |
have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)" |
275 |
using Q measurable_pair_swap' by (auto intro: measurable_sets) |
|
276 |
note Q.measurable_emeasure_Pair1[OF this] |
|
277 |
moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^isub>M M1)) = (\<lambda>x. (x, y)) -` Q" |
|
278 |
using Q[THEN sets_into_space] by (auto simp: space_pair_measure) |
|
279 |
ultimately show ?thesis by simp |
|
39088
ca17017c10e6
Measurable on product space is equiv. to measurable components
hoelzl
parents:
39082
diff
changeset
|
280 |
qed |
ca17017c10e6
Measurable on product space is equiv. to measurable components
hoelzl
parents:
39082
diff
changeset
|
281 |
|
47694 | 282 |
lemma (in pair_sigma_finite) emeasure_pair_measure: |
283 |
assumes "X \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
|
284 |
shows "emeasure (M1 \<Otimes>\<^isub>M M2) X = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator X (x, y) \<partial>M2 \<partial>M1)" (is "_ = ?\<mu> X") |
|
285 |
proof (rule emeasure_measure_of[OF pair_measure_def]) |
|
286 |
show "positive (sets (M1 \<Otimes>\<^isub>M M2)) ?\<mu>" |
|
287 |
by (auto simp: positive_def positive_integral_positive) |
|
288 |
have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y" |
|
289 |
by (auto simp: indicator_def) |
|
290 |
show "countably_additive (sets (M1 \<Otimes>\<^isub>M M2)) ?\<mu>" |
|
291 |
proof (rule countably_additiveI) |
|
292 |
fix F :: "nat \<Rightarrow> ('a \<times> 'b) set" assume F: "range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2)" "disjoint_family F" |
|
293 |
from F have *: "\<And>i. F i \<in> sets (M1 \<Otimes>\<^isub>M M2)" "(\<Union>i. F i) \<in> sets (M1 \<Otimes>\<^isub>M M2)" by auto |
|
294 |
moreover from F have "\<And>i. (\<lambda>x. emeasure M2 (Pair x -` F i)) \<in> borel_measurable M1" |
|
295 |
by (intro measurable_emeasure_Pair1) auto |
|
296 |
moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)" |
|
297 |
by (intro disjoint_family_on_bisimulation[OF F(2)]) auto |
|
298 |
moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M2" |
|
299 |
using F by (auto simp: sets_Pair1) |
|
300 |
ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)" |
|
301 |
by (auto simp add: vimage_UN positive_integral_suminf[symmetric] suminf_emeasure subset_eq emeasure_nonneg sets_Pair1 |
|
302 |
intro!: positive_integral_cong positive_integral_indicator[symmetric]) |
|
303 |
qed |
|
304 |
show "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2} \<subseteq> Pow (space M1 \<times> space M2)" |
|
305 |
using space_closed[of M1] space_closed[of M2] by auto |
|
306 |
qed fact |
|
307 |
||
308 |
lemma (in pair_sigma_finite) emeasure_pair_measure_alt: |
|
309 |
assumes X: "X \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
|
310 |
shows "emeasure (M1 \<Otimes>\<^isub>M M2) X = (\<integral>\<^isup>+x. emeasure M2 (Pair x -` X) \<partial>M1)" |
|
311 |
proof - |
|
312 |
have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y" |
|
313 |
by (auto simp: indicator_def) |
|
314 |
show ?thesis |
|
315 |
using X by (auto intro!: positive_integral_cong simp: emeasure_pair_measure sets_Pair1) |
|
40859 | 316 |
qed |
317 |
||
47694 | 318 |
lemma (in pair_sigma_finite) emeasure_pair_measure_Times: |
319 |
assumes A: "A \<in> sets M1" and B: "B \<in> sets M2" |
|
320 |
shows "emeasure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = emeasure M1 A * emeasure M2 B" |
|
40859 | 321 |
proof - |
47694 | 322 |
have "emeasure (M1 \<Otimes>\<^isub>M M2) (A \<times> B) = (\<integral>\<^isup>+x. emeasure M2 B * indicator A x \<partial>M1)" |
323 |
using A B by (auto intro!: positive_integral_cong simp: emeasure_pair_measure_alt) |
|
324 |
also have "\<dots> = emeasure M2 B * emeasure M1 A" |
|
325 |
using A by (simp add: emeasure_nonneg positive_integral_cmult_indicator) |
|
326 |
finally show ?thesis |
|
327 |
by (simp add: ac_simps) |
|
40859 | 328 |
qed |
329 |
||
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
330 |
lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator: |
47694 | 331 |
defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}" |
332 |
shows "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> E \<and> incseq F \<and> (\<Union>i. F i) = space M1 \<times> space M2 \<and> |
|
333 |
(\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)" |
|
40859 | 334 |
proof - |
47694 | 335 |
from M1.sigma_finite_incseq guess F1 . note F1 = this |
336 |
from M2.sigma_finite_incseq guess F2 . note F2 = this |
|
337 |
from F1 F2 have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto |
|
40859 | 338 |
let ?F = "\<lambda>i. F1 i \<times> F2 i" |
47694 | 339 |
show ?thesis |
40859 | 340 |
proof (intro exI[of _ ?F] conjI allI) |
47694 | 341 |
show "range ?F \<subseteq> E" using F1 F2 by (auto simp: E_def) (metis range_subsetD) |
40859 | 342 |
next |
343 |
have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)" |
|
344 |
proof (intro subsetI) |
|
345 |
fix x assume "x \<in> space M1 \<times> space M2" |
|
346 |
then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j" |
|
347 |
by (auto simp: space) |
|
348 |
then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
349 |
using `incseq F1` `incseq F2` unfolding incseq_def |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
350 |
by (force split: split_max)+ |
40859 | 351 |
then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)" |
352 |
by (intro SigmaI) (auto simp add: min_max.sup_commute) |
|
353 |
then show "x \<in> (\<Union>i. ?F i)" by auto |
|
354 |
qed |
|
47694 | 355 |
then show "(\<Union>i. ?F i) = space M1 \<times> space M2" |
356 |
using space by (auto simp: space) |
|
40859 | 357 |
next |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
358 |
fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
359 |
using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto |
40859 | 360 |
next |
361 |
fix i |
|
362 |
from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto |
|
47694 | 363 |
with F1 F2 emeasure_nonneg[of M1 "F1 i"] emeasure_nonneg[of M2 "F2 i"] |
364 |
show "emeasure (M1 \<Otimes>\<^isub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>" |
|
365 |
by (auto simp add: emeasure_pair_measure_Times) |
|
366 |
qed |
|
367 |
qed |
|
368 |
||
369 |
sublocale pair_sigma_finite \<subseteq> sigma_finite_measure "M1 \<Otimes>\<^isub>M M2" |
|
370 |
proof |
|
371 |
from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this |
|
372 |
show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2) \<and> (\<Union>i. F i) = space (M1 \<Otimes>\<^isub>M M2) \<and> (\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>)" |
|
373 |
proof (rule exI[of _ F], intro conjI) |
|
374 |
show "range F \<subseteq> sets (M1 \<Otimes>\<^isub>M M2)" using F by (auto simp: pair_measure_def) |
|
375 |
show "(\<Union>i. F i) = space (M1 \<Otimes>\<^isub>M M2)" |
|
376 |
using F by (auto simp: space_pair_measure) |
|
377 |
show "\<forall>i. emeasure (M1 \<Otimes>\<^isub>M M2) (F i) \<noteq> \<infinity>" using F by auto |
|
40859 | 378 |
qed |
379 |
qed |
|
380 |
||
47694 | 381 |
lemma sigma_finite_pair_measure: |
382 |
assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B" |
|
383 |
shows "sigma_finite_measure (A \<Otimes>\<^isub>M B)" |
|
384 |
proof - |
|
385 |
interpret A: sigma_finite_measure A by fact |
|
386 |
interpret B: sigma_finite_measure B by fact |
|
387 |
interpret AB: pair_sigma_finite A B .. |
|
388 |
show ?thesis .. |
|
40859 | 389 |
qed |
39088
ca17017c10e6
Measurable on product space is equiv. to measurable components
hoelzl
parents:
39082
diff
changeset
|
390 |
|
47694 | 391 |
lemma sets_pair_swap: |
392 |
assumes "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
393 |
shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1) \<in> sets (M2 \<Otimes>\<^isub>M M1)" |
47694 | 394 |
using measurable_pair_swap' assms by (rule measurable_sets) |
41661 | 395 |
|
47694 | 396 |
lemma (in pair_sigma_finite) distr_pair_swap: |
397 |
"M1 \<Otimes>\<^isub>M M2 = distr (M2 \<Otimes>\<^isub>M M1) (M1 \<Otimes>\<^isub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D") |
|
40859 | 398 |
proof - |
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
399 |
interpret Q: pair_sigma_finite M2 M1 by default |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
400 |
from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this |
47694 | 401 |
let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}" |
402 |
show ?thesis |
|
403 |
proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]]) |
|
404 |
show "?E \<subseteq> Pow (space ?P)" |
|
405 |
using space_closed[of M1] space_closed[of M2] by (auto simp: space_pair_measure) |
|
406 |
show "sets ?P = sigma_sets (space ?P) ?E" |
|
407 |
by (simp add: sets_pair_measure space_pair_measure) |
|
408 |
then show "sets ?D = sigma_sets (space ?P) ?E" |
|
409 |
by simp |
|
410 |
next |
|
411 |
show "range F \<subseteq> ?E" "incseq F" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>" |
|
412 |
using F by (auto simp: space_pair_measure) |
|
413 |
next |
|
414 |
fix X assume "X \<in> ?E" |
|
415 |
then obtain A B where X[simp]: "X = A \<times> B" and A: "A \<in> sets M1" and B: "B \<in> sets M2" by auto |
|
416 |
have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^isub>M M1) = B \<times> A" |
|
417 |
using sets_into_space[OF A] sets_into_space[OF B] by (auto simp: space_pair_measure) |
|
418 |
with A B show "emeasure (M1 \<Otimes>\<^isub>M M2) X = emeasure ?D X" |
|
419 |
by (simp add: emeasure_pair_measure_Times Q.emeasure_pair_measure_Times emeasure_distr |
|
420 |
measurable_pair_swap' ac_simps) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
421 |
qed |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
422 |
qed |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
423 |
|
47694 | 424 |
lemma (in pair_sigma_finite) emeasure_pair_measure_alt2: |
425 |
assumes A: "A \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
|
426 |
shows "emeasure (M1 \<Otimes>\<^isub>M M2) A = (\<integral>\<^isup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)" |
|
427 |
(is "_ = ?\<nu> A") |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
428 |
proof - |
47694 | 429 |
interpret Q: pair_sigma_finite M2 M1 by default |
430 |
have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^isub>M M1))) = (\<lambda>x. (x, y)) -` A" |
|
431 |
using sets_into_space[OF A] by (auto simp: space_pair_measure) |
|
432 |
show ?thesis using A |
|
433 |
by (subst distr_pair_swap) |
|
434 |
(simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap'] |
|
435 |
Q.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A]) |
|
40859 | 436 |
qed |
437 |
||
438 |
section "Fubinis theorem" |
|
439 |
||
440 |
lemma (in pair_sigma_finite) simple_function_cut: |
|
47694 | 441 |
assumes f: "simple_function (M1 \<Otimes>\<^isub>M M2) f" "\<And>x. 0 \<le> f x" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
442 |
shows "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1" |
47694 | 443 |
and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" |
40859 | 444 |
proof - |
47694 | 445 |
have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
446 |
using f(1) by (rule borel_measurable_simple_function) |
47694 | 447 |
let ?F = "\<lambda>z. f -` {z} \<inter> space (M1 \<Otimes>\<^isub>M M2)" |
46731 | 448 |
let ?F' = "\<lambda>x z. Pair x -` ?F z" |
40859 | 449 |
{ fix x assume "x \<in> space M1" |
450 |
have [simp]: "\<And>z y. indicator (?F z) (x, y) = indicator (?F' x z) y" |
|
451 |
by (auto simp: indicator_def) |
|
47694 | 452 |
have "\<And>y. y \<in> space M2 \<Longrightarrow> (x, y) \<in> space (M1 \<Otimes>\<^isub>M M2)" using `x \<in> space M1` |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
453 |
by (simp add: space_pair_measure) |
40859 | 454 |
moreover have "\<And>x z. ?F' x z \<in> sets M2" using f_borel |
47694 | 455 |
by (rule sets_Pair1[OF measurable_sets]) auto |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
456 |
ultimately have "simple_function M2 (\<lambda> y. f (x, y))" |
47694 | 457 |
apply (rule_tac simple_function_cong[THEN iffD2, OF _]) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
458 |
apply (rule simple_function_indicator_representation[OF f(1)]) |
47694 | 459 |
using `x \<in> space M1` by auto } |
40859 | 460 |
note M2_sf = this |
461 |
{ fix x assume x: "x \<in> space M1" |
|
47694 | 462 |
then have "(\<integral>\<^isup>+y. f (x, y) \<partial>M2) = (\<Sum>z\<in>f ` space (M1 \<Otimes>\<^isub>M M2). z * emeasure M2 (?F' x z))" |
463 |
unfolding positive_integral_eq_simple_integral[OF M2_sf[OF x] f(2)] |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
464 |
unfolding simple_integral_def |
40859 | 465 |
proof (safe intro!: setsum_mono_zero_cong_left) |
47694 | 466 |
from f(1) show "finite (f ` space (M1 \<Otimes>\<^isub>M M2))" by (rule simple_functionD) |
40859 | 467 |
next |
47694 | 468 |
fix y assume "y \<in> space M2" then show "f (x, y) \<in> f ` space (M1 \<Otimes>\<^isub>M M2)" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
469 |
using `x \<in> space M1` by (auto simp: space_pair_measure) |
40859 | 470 |
next |
47694 | 471 |
fix x' y assume "(x', y) \<in> space (M1 \<Otimes>\<^isub>M M2)" |
40859 | 472 |
"f (x', y) \<notin> (\<lambda>y. f (x, y)) ` space M2" |
473 |
then have *: "?F' x (f (x', y)) = {}" |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
474 |
by (force simp: space_pair_measure) |
47694 | 475 |
show "f (x', y) * emeasure M2 (?F' x (f (x', y))) = 0" |
40859 | 476 |
unfolding * by simp |
477 |
qed (simp add: vimage_compose[symmetric] comp_def |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
478 |
space_pair_measure) } |
40859 | 479 |
note eq = this |
47694 | 480 |
moreover have "\<And>z. ?F z \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
481 |
by (auto intro!: f_borel borel_measurable_vimage) |
|
482 |
moreover then have "\<And>z. (\<lambda>x. emeasure M2 (?F' x z)) \<in> borel_measurable M1" |
|
483 |
by (auto intro!: measurable_emeasure_Pair1 simp del: vimage_Int) |
|
484 |
moreover have *: "\<And>i x. 0 \<le> emeasure M2 (Pair x -` (f -` {i} \<inter> space (M1 \<Otimes>\<^isub>M M2)))" |
|
485 |
using f(1)[THEN simple_functionD(2)] f(2) by (intro emeasure_nonneg) |
|
486 |
moreover { fix i assume "i \<in> f`space (M1 \<Otimes>\<^isub>M M2)" |
|
487 |
with * have "\<And>x. 0 \<le> i * emeasure M2 (Pair x -` (f -` {i} \<inter> space (M1 \<Otimes>\<^isub>M M2)))" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
488 |
using f(2) by auto } |
40859 | 489 |
ultimately |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
490 |
show "(\<lambda>x. \<integral>\<^isup>+y. f (x, y) \<partial>M2) \<in> borel_measurable M1" |
47694 | 491 |
and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" using f(2) |
492 |
by (auto simp del: vimage_Int cong: measurable_cong intro!: setsum_cong |
|
493 |
simp add: positive_integral_setsum simple_integral_def |
|
494 |
positive_integral_cmult |
|
495 |
positive_integral_cong[OF eq] |
|
40859 | 496 |
positive_integral_eq_simple_integral[OF f] |
47694 | 497 |
emeasure_pair_measure_alt[symmetric]) |
40859 | 498 |
qed |
499 |
||
500 |
lemma (in pair_sigma_finite) positive_integral_fst_measurable: |
|
47694 | 501 |
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
502 |
shows "(\<lambda>x. \<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<in> borel_measurable M1" |
40859 | 503 |
(is "?C f \<in> borel_measurable M1") |
47694 | 504 |
and "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" |
40859 | 505 |
proof - |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
506 |
from borel_measurable_implies_simple_function_sequence'[OF f] guess F . note F = this |
47694 | 507 |
then have F_borel: "\<And>i. F i \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
40859 | 508 |
by (auto intro: borel_measurable_simple_function) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
509 |
note sf = simple_function_cut[OF F(1,5)] |
41097
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents:
41096
diff
changeset
|
510 |
then have "(\<lambda>x. SUP i. ?C (F i) x) \<in> borel_measurable M1" |
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents:
41096
diff
changeset
|
511 |
using F(1) by auto |
40859 | 512 |
moreover |
513 |
{ fix x assume "x \<in> space M1" |
|
47694 | 514 |
from F measurable_Pair2[OF F_borel `x \<in> space M1`] |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
515 |
have "(\<integral>\<^isup>+y. (SUP i. F i (x, y)) \<partial>M2) = (SUP i. ?C (F i) x)" |
47694 | 516 |
by (intro positive_integral_monotone_convergence_SUP) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
517 |
(auto simp: incseq_Suc_iff le_fun_def) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
518 |
then have "(SUP i. ?C (F i) x) = ?C f x" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
519 |
unfolding F(4) positive_integral_max_0 by simp } |
40859 | 520 |
note SUPR_C = this |
521 |
ultimately show "?C f \<in> borel_measurable M1" |
|
41097
a1abfa4e2b44
use SUPR_ and INFI_apply instead of SUPR_, INFI_fun_expand
hoelzl
parents:
41096
diff
changeset
|
522 |
by (simp cong: measurable_cong) |
47694 | 523 |
have "(\<integral>\<^isup>+x. (SUP i. F i x) \<partial>(M1 \<Otimes>\<^isub>M M2)) = (SUP i. integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (F i))" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
524 |
using F_borel F |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
525 |
by (intro positive_integral_monotone_convergence_SUP) auto |
47694 | 526 |
also have "(SUP i. integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (F i)) = (SUP i. \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1)" |
40859 | 527 |
unfolding sf(2) by simp |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
528 |
also have "\<dots> = \<integral>\<^isup>+ x. (SUP i. \<integral>\<^isup>+ y. F i (x, y) \<partial>M2) \<partial>M1" using F sf(1) |
47694 | 529 |
by (intro positive_integral_monotone_convergence_SUP[symmetric]) |
530 |
(auto intro!: positive_integral_mono positive_integral_positive |
|
531 |
simp: incseq_Suc_iff le_fun_def) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
532 |
also have "\<dots> = \<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. (SUP i. F i (x, y)) \<partial>M2) \<partial>M1" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
533 |
using F_borel F(2,5) |
47694 | 534 |
by (auto intro!: positive_integral_cong positive_integral_monotone_convergence_SUP[symmetric] measurable_Pair2 |
535 |
simp: incseq_Suc_iff le_fun_def) |
|
536 |
finally show "(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
537 |
using F by (simp add: positive_integral_max_0) |
40859 | 538 |
qed |
539 |
||
47694 | 540 |
lemma (in pair_sigma_finite) positive_integral_snd_measurable: |
541 |
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
|
542 |
shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" |
|
41661 | 543 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
544 |
interpret Q: pair_sigma_finite M2 M1 by default |
47694 | 545 |
note measurable_pair_swap[OF f] |
40859 | 546 |
from Q.positive_integral_fst_measurable[OF this] |
47694 | 547 |
have "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1))" |
40859 | 548 |
by simp |
47694 | 549 |
also have "(\<integral>\<^isup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) f" |
550 |
by (subst distr_pair_swap) |
|
551 |
(auto simp: positive_integral_distr[OF measurable_pair_swap' f] intro!: positive_integral_cong) |
|
40859 | 552 |
finally show ?thesis . |
553 |
qed |
|
554 |
||
555 |
lemma (in pair_sigma_finite) Fubini: |
|
47694 | 556 |
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
557 |
shows "(\<integral>\<^isup>+ y. (\<integral>\<^isup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x, y) \<partial>M2) \<partial>M1)" |
40859 | 558 |
unfolding positive_integral_snd_measurable[OF assms] |
559 |
unfolding positive_integral_fst_measurable[OF assms] .. |
|
560 |
||
561 |
lemma (in pair_sigma_finite) AE_pair: |
|
47694 | 562 |
assumes "AE x in (M1 \<Otimes>\<^isub>M M2). Q x" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
563 |
shows "AE x in M1. (AE y in M2. Q (x, y))" |
40859 | 564 |
proof - |
47694 | 565 |
obtain N where N: "N \<in> sets (M1 \<Otimes>\<^isub>M M2)" "emeasure (M1 \<Otimes>\<^isub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^isub>M M2). \<not> Q x} \<subseteq> N" |
566 |
using assms unfolding eventually_ae_filter by auto |
|
40859 | 567 |
show ?thesis |
47694 | 568 |
proof (rule AE_I) |
569 |
from N measurable_emeasure_Pair1[OF `N \<in> sets (M1 \<Otimes>\<^isub>M M2)`] |
|
570 |
show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0" |
|
571 |
by (auto simp: emeasure_pair_measure_alt positive_integral_0_iff emeasure_nonneg) |
|
572 |
show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1" |
|
573 |
by (intro borel_measurable_ereal_neq_const measurable_emeasure_Pair1 N) |
|
574 |
{ fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0" |
|
575 |
have "AE y in M2. Q (x, y)" |
|
576 |
proof (rule AE_I) |
|
577 |
show "emeasure M2 (Pair x -` N) = 0" by fact |
|
578 |
show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1) |
|
40859 | 579 |
show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N" |
47694 | 580 |
using N `x \<in> space M1` unfolding space_pair_measure by auto |
40859 | 581 |
qed } |
47694 | 582 |
then show "{x \<in> space M1. \<not> (AE y in M2. Q (x, y))} \<subseteq> {x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0}" |
40859 | 583 |
by auto |
39088
ca17017c10e6
Measurable on product space is equiv. to measurable components
hoelzl
parents:
39082
diff
changeset
|
584 |
qed |
ca17017c10e6
Measurable on product space is equiv. to measurable components
hoelzl
parents:
39082
diff
changeset
|
585 |
qed |
35833 | 586 |
|
47694 | 587 |
lemma (in pair_sigma_finite) AE_pair_measure: |
588 |
assumes "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P x} \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
|
589 |
assumes ae: "AE x in M1. AE y in M2. P (x, y)" |
|
590 |
shows "AE x in M1 \<Otimes>\<^isub>M M2. P x" |
|
591 |
proof (subst AE_iff_measurable[OF _ refl]) |
|
592 |
show "{x\<in>space (M1 \<Otimes>\<^isub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
|
593 |
by (rule sets_Collect) fact |
|
594 |
then have "emeasure (M1 \<Otimes>\<^isub>M M2) {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} = |
|
595 |
(\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)" |
|
596 |
by (simp add: emeasure_pair_measure) |
|
597 |
also have "\<dots> = (\<integral>\<^isup>+ x. \<integral>\<^isup>+ y. 0 \<partial>M2 \<partial>M1)" |
|
598 |
using ae |
|
599 |
apply (safe intro!: positive_integral_cong_AE) |
|
600 |
apply (intro AE_I2) |
|
601 |
apply (safe intro!: positive_integral_cong_AE) |
|
602 |
apply auto |
|
603 |
done |
|
604 |
finally show "emeasure (M1 \<Otimes>\<^isub>M M2) {x \<in> space (M1 \<Otimes>\<^isub>M M2). \<not> P x} = 0" by simp |
|
605 |
qed |
|
606 |
||
607 |
lemma (in pair_sigma_finite) AE_pair_iff: |
|
608 |
"{x\<in>space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow> |
|
609 |
(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x))" |
|
610 |
using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto |
|
611 |
||
612 |
lemma AE_distr_iff: |
|
613 |
assumes f: "f \<in> measurable M N" and P: "{x \<in> space N. P x} \<in> sets N" |
|
614 |
shows "(AE x in distr M N f. P x) \<longleftrightarrow> (AE x in M. P (f x))" |
|
615 |
proof (subst (1 2) AE_iff_measurable[OF _ refl]) |
|
616 |
from P show "{x \<in> space (distr M N f). \<not> P x} \<in> sets (distr M N f)" |
|
617 |
by (auto intro!: sets_Collect_neg) |
|
618 |
moreover |
|
619 |
have "f -` {x \<in> space N. P x} \<inter> space M = {x \<in> space M. P (f x)}" |
|
620 |
using f by (auto dest: measurable_space) |
|
621 |
then show "{x \<in> space M. \<not> P (f x)} \<in> sets M" |
|
622 |
using measurable_sets[OF f P] by (auto intro!: sets_Collect_neg) |
|
623 |
moreover have "f -` {x\<in>space N. \<not> P x} \<inter> space M = {x \<in> space M. \<not> P (f x)}" |
|
624 |
using f by (auto dest: measurable_space) |
|
625 |
ultimately show "(emeasure (distr M N f) {x \<in> space (distr M N f). \<not> P x} = 0) = |
|
626 |
(emeasure M {x \<in> space M. \<not> P (f x)} = 0)" |
|
627 |
using f by (simp add: emeasure_distr) |
|
628 |
qed |
|
629 |
||
630 |
lemma (in pair_sigma_finite) AE_commute: |
|
631 |
assumes P: "{x\<in>space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^isub>M M2)" |
|
632 |
shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)" |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
633 |
proof - |
47694 | 634 |
interpret Q: pair_sigma_finite M2 M1 .. |
635 |
have [simp]: "\<And>x. (fst (case x of (x, y) \<Rightarrow> (y, x))) = snd x" "\<And>x. (snd (case x of (x, y) \<Rightarrow> (y, x))) = fst x" |
|
636 |
by auto |
|
637 |
have "{x \<in> space (M2 \<Otimes>\<^isub>M M1). P (snd x) (fst x)} = |
|
638 |
(\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^isub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^isub>M M1)" |
|
639 |
by (auto simp: space_pair_measure) |
|
640 |
also have "\<dots> \<in> sets (M2 \<Otimes>\<^isub>M M1)" |
|
641 |
by (intro sets_pair_swap P) |
|
642 |
finally show ?thesis |
|
643 |
apply (subst AE_pair_iff[OF P]) |
|
644 |
apply (subst distr_pair_swap) |
|
645 |
apply (subst AE_distr_iff[OF measurable_pair_swap' P]) |
|
646 |
apply (subst Q.AE_pair_iff) |
|
647 |
apply simp_all |
|
648 |
done |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
649 |
qed |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
650 |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
651 |
lemma (in pair_sigma_finite) integrable_product_swap: |
47694 | 652 |
assumes "integrable (M1 \<Otimes>\<^isub>M M2) f" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
653 |
shows "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x))" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
654 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
655 |
interpret Q: pair_sigma_finite M2 M1 by default |
41661 | 656 |
have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff) |
657 |
show ?thesis unfolding * |
|
47694 | 658 |
by (rule integrable_distr[OF measurable_pair_swap']) |
659 |
(simp add: distr_pair_swap[symmetric] assms) |
|
41661 | 660 |
qed |
661 |
||
662 |
lemma (in pair_sigma_finite) integrable_product_swap_iff: |
|
47694 | 663 |
"integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^isub>M M2) f" |
41661 | 664 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
665 |
interpret Q: pair_sigma_finite M2 M1 by default |
41661 | 666 |
from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f] |
667 |
show ?thesis by auto |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
668 |
qed |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
669 |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
670 |
lemma (in pair_sigma_finite) integral_product_swap: |
47694 | 671 |
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
672 |
shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^isub>M M1)) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f" |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
673 |
proof - |
41661 | 674 |
have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff) |
47694 | 675 |
show ?thesis unfolding * |
676 |
by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric]) |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
677 |
qed |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
678 |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
679 |
lemma (in pair_sigma_finite) integrable_fst_measurable: |
47694 | 680 |
assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f" |
681 |
shows "AE x in M1. integrable M2 (\<lambda> y. f (x, y))" (is "?AE") |
|
682 |
and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f" (is "?INT") |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
683 |
proof - |
47694 | 684 |
have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
685 |
using f by auto |
|
46731 | 686 |
let ?pf = "\<lambda>x. ereal (f x)" and ?nf = "\<lambda>x. ereal (- f x)" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
687 |
have |
47694 | 688 |
borel: "?nf \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)""?pf \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" and |
689 |
int: "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) ?nf \<noteq> \<infinity>" "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) ?pf \<noteq> \<infinity>" |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
690 |
using assms by auto |
43920 | 691 |
have "(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>" |
692 |
"(\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>" |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
693 |
using borel[THEN positive_integral_fst_measurable(1)] int |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
694 |
unfolding borel[THEN positive_integral_fst_measurable(2)] by simp_all |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
695 |
with borel[THEN positive_integral_fst_measurable(1)] |
43920 | 696 |
have AE_pos: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>" |
697 |
"AE x in M1. (\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>" |
|
47694 | 698 |
by (auto intro!: positive_integral_PInf_AE ) |
43920 | 699 |
then have AE: "AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>" |
700 |
"AE x in M1. \<bar>\<integral>\<^isup>+y. ereal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>" |
|
47694 | 701 |
by (auto simp: positive_integral_positive) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
702 |
from AE_pos show ?AE using assms |
47694 | 703 |
by (simp add: measurable_Pair2[OF f_borel] integrable_def) |
43920 | 704 |
{ fix f have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)" |
47694 | 705 |
using positive_integral_positive |
706 |
by (intro positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder) |
|
43920 | 707 |
then have "(\<integral>\<^isup>+ x. - \<integral>\<^isup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = 0" by simp } |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
708 |
note this[simp] |
47694 | 709 |
{ fix f assume borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
710 |
and int: "integral\<^isup>P (M1 \<Otimes>\<^isub>M M2) (\<lambda>x. ereal (f x)) \<noteq> \<infinity>" |
|
711 |
and AE: "AE x in M1. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>" |
|
43920 | 712 |
have "integrable M1 (\<lambda>x. real (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f") |
41705 | 713 |
proof (intro integrable_def[THEN iffD2] conjI) |
714 |
show "?f \<in> borel_measurable M1" |
|
47694 | 715 |
using borel by (auto intro!: positive_integral_fst_measurable) |
43920 | 716 |
have "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) = (\<integral>\<^isup>+x. (\<integral>\<^isup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1)" |
47694 | 717 |
using AE positive_integral_positive[of M2] |
718 |
by (auto intro!: positive_integral_cong_AE simp: ereal_real) |
|
43920 | 719 |
then show "(\<integral>\<^isup>+x. ereal (?f x) \<partial>M1) \<noteq> \<infinity>" |
41705 | 720 |
using positive_integral_fst_measurable[OF borel] int by simp |
43920 | 721 |
have "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) = (\<integral>\<^isup>+x. 0 \<partial>M1)" |
47694 | 722 |
by (intro positive_integral_cong_pos) |
723 |
(simp add: positive_integral_positive real_of_ereal_pos) |
|
43920 | 724 |
then show "(\<integral>\<^isup>+x. ereal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp |
41705 | 725 |
qed } |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
726 |
with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)] |
41705 | 727 |
show ?INT |
47694 | 728 |
unfolding lebesgue_integral_def[of "M1 \<Otimes>\<^isub>M M2"] lebesgue_integral_def[of M2] |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
729 |
borel[THEN positive_integral_fst_measurable(2), symmetric] |
47694 | 730 |
using AE[THEN integral_real] |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
731 |
by simp |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
732 |
qed |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
733 |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
734 |
lemma (in pair_sigma_finite) integrable_snd_measurable: |
47694 | 735 |
assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f" |
736 |
shows "AE y in M2. integrable M1 (\<lambda>x. f (x, y))" (is "?AE") |
|
737 |
and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^isup>L (M1 \<Otimes>\<^isub>M M2) f" (is "?INT") |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
738 |
proof - |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
739 |
interpret Q: pair_sigma_finite M2 M1 by default |
47694 | 740 |
have Q_int: "integrable (M2 \<Otimes>\<^isub>M M1) (\<lambda>(x, y). f (y, x))" |
41661 | 741 |
using f unfolding integrable_product_swap_iff . |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
742 |
show ?INT |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
743 |
using Q.integrable_fst_measurable(2)[OF Q_int] |
47694 | 744 |
using integral_product_swap[of f] f by auto |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
745 |
show ?AE |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
746 |
using Q.integrable_fst_measurable(1)[OF Q_int] |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
747 |
by simp |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
748 |
qed |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
749 |
|
47694 | 750 |
lemma (in pair_sigma_finite) positive_integral_fst_measurable': |
751 |
assumes f: "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
|
752 |
shows "(\<lambda>x. \<integral>\<^isup>+ y. f x y \<partial>M2) \<in> borel_measurable M1" |
|
753 |
using positive_integral_fst_measurable(1)[OF f] by simp |
|
754 |
||
755 |
lemma (in pair_sigma_finite) integral_fst_measurable: |
|
756 |
"(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow> (\<lambda>x. \<integral> y. f x y \<partial>M2) \<in> borel_measurable M1" |
|
757 |
by (auto simp: lebesgue_integral_def intro!: borel_measurable_diff positive_integral_fst_measurable') |
|
758 |
||
759 |
lemma (in pair_sigma_finite) positive_integral_snd_measurable': |
|
760 |
assumes f: "(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
|
761 |
shows "(\<lambda>y. \<integral>\<^isup>+ x. f x y \<partial>M1) \<in> borel_measurable M2" |
|
762 |
proof - |
|
763 |
interpret Q: pair_sigma_finite M2 M1 .. |
|
764 |
show ?thesis |
|
765 |
using measurable_pair_swap[OF f] |
|
766 |
by (intro Q.positive_integral_fst_measurable') (simp add: split_beta') |
|
767 |
qed |
|
768 |
||
769 |
lemma (in pair_sigma_finite) integral_snd_measurable: |
|
770 |
"(\<lambda>x. f (fst x) (snd x)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2) \<Longrightarrow> (\<lambda>y. \<integral> x. f x y \<partial>M1) \<in> borel_measurable M2" |
|
771 |
by (auto simp: lebesgue_integral_def intro!: borel_measurable_diff positive_integral_snd_measurable') |
|
772 |
||
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
773 |
lemma (in pair_sigma_finite) Fubini_integral: |
47694 | 774 |
assumes f: "integrable (M1 \<Otimes>\<^isub>M M2) f" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
775 |
shows "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1)" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
776 |
unfolding integrable_snd_measurable[OF assms] |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
777 |
unfolding integrable_fst_measurable[OF assms] .. |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
778 |
|
47694 | 779 |
section {* Products on counting spaces, densities and distributions *} |
40859 | 780 |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
781 |
lemma sigma_sets_pair_measure_generator_finite: |
38656 | 782 |
assumes "finite A" and "finite B" |
47694 | 783 |
shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<subseteq> A \<and> b \<subseteq> B} = Pow (A \<times> B)" |
40859 | 784 |
(is "sigma_sets ?prod ?sets = _") |
38656 | 785 |
proof safe |
786 |
have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product) |
|
787 |
fix x assume subset: "x \<subseteq> A \<times> B" |
|
788 |
hence "finite x" using fin by (rule finite_subset) |
|
40859 | 789 |
from this subset show "x \<in> sigma_sets ?prod ?sets" |
38656 | 790 |
proof (induct x) |
791 |
case empty show ?case by (rule sigma_sets.Empty) |
|
792 |
next |
|
793 |
case (insert a x) |
|
47694 | 794 |
hence "{a} \<in> sigma_sets ?prod ?sets" by auto |
38656 | 795 |
moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto |
796 |
ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un) |
|
797 |
qed |
|
798 |
next |
|
799 |
fix x a b |
|
40859 | 800 |
assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x" |
38656 | 801 |
from sigma_sets_into_sp[OF _ this(1)] this(2) |
40859 | 802 |
show "a \<in> A" and "b \<in> B" by auto |
35833 | 803 |
qed |
804 |
||
47694 | 805 |
lemma pair_measure_count_space: |
806 |
assumes A: "finite A" and B: "finite B" |
|
807 |
shows "count_space A \<Otimes>\<^isub>M count_space B = count_space (A \<times> B)" (is "?P = ?C") |
|
808 |
proof (rule measure_eqI) |
|
809 |
interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact |
|
810 |
interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact |
|
811 |
interpret P: pair_sigma_finite "count_space A" "count_space B" by default |
|
812 |
show eq: "sets ?P = sets ?C" |
|
813 |
by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B) |
|
814 |
fix X assume X: "X \<in> sets ?P" |
|
815 |
with eq have X_subset: "X \<subseteq> A \<times> B" by simp |
|
816 |
with A B have fin_Pair: "\<And>x. finite (Pair x -` X)" |
|
817 |
by (intro finite_subset[OF _ B]) auto |
|
818 |
have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B) |
|
819 |
show "emeasure ?P X = emeasure ?C X" |
|
820 |
apply (subst P.emeasure_pair_measure_alt[OF X]) |
|
821 |
apply (subst emeasure_count_space) |
|
822 |
using X_subset apply auto [] |
|
823 |
apply (simp add: fin_Pair emeasure_count_space X_subset fin_X) |
|
824 |
apply (subst positive_integral_count_space) |
|
825 |
using A apply simp |
|
826 |
apply (simp del: real_of_nat_setsum add: real_of_nat_setsum[symmetric]) |
|
827 |
apply (subst card_gt_0_iff) |
|
828 |
apply (simp add: fin_Pair) |
|
829 |
apply (subst card_SigmaI[symmetric]) |
|
830 |
using A apply simp |
|
831 |
using fin_Pair apply simp |
|
832 |
using X_subset apply (auto intro!: arg_cong[where f=card]) |
|
833 |
done |
|
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:
44890
diff
changeset
|
834 |
qed |
35833 | 835 |
|
47694 | 836 |
lemma pair_measure_density: |
837 |
assumes f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x" |
|
838 |
assumes g: "g \<in> borel_measurable M2" "AE x in M2. 0 \<le> g x" |
|
839 |
assumes "sigma_finite_measure M1" "sigma_finite_measure M2" |
|
840 |
assumes "sigma_finite_measure (density M1 f)" "sigma_finite_measure (density M2 g)" |
|
841 |
shows "density M1 f \<Otimes>\<^isub>M density M2 g = density (M1 \<Otimes>\<^isub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R") |
|
842 |
proof (rule measure_eqI) |
|
843 |
interpret M1: sigma_finite_measure M1 by fact |
|
844 |
interpret M2: sigma_finite_measure M2 by fact |
|
845 |
interpret D1: sigma_finite_measure "density M1 f" by fact |
|
846 |
interpret D2: sigma_finite_measure "density M2 g" by fact |
|
847 |
interpret L: pair_sigma_finite "density M1 f" "density M2 g" .. |
|
848 |
interpret R: pair_sigma_finite M1 M2 .. |
|
849 |
||
850 |
fix A assume A: "A \<in> sets ?L" |
|
851 |
then have indicator_eq: "\<And>x y. indicator A (x, y) = indicator (Pair x -` A) y" |
|
852 |
and Pair_A: "\<And>x. Pair x -` A \<in> sets M2" |
|
853 |
by (auto simp: indicator_def sets_Pair1) |
|
854 |
have f_fst: "(\<lambda>p. f (fst p)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
|
855 |
using measurable_comp[OF measurable_fst f(1)] by (simp add: comp_def) |
|
856 |
have g_snd: "(\<lambda>p. g (snd p)) \<in> borel_measurable (M1 \<Otimes>\<^isub>M M2)" |
|
857 |
using measurable_comp[OF measurable_snd g(1)] by (simp add: comp_def) |
|
858 |
have "(\<lambda>x. \<integral>\<^isup>+ y. g (snd (x, y)) * indicator A (x, y) \<partial>M2) \<in> borel_measurable M1" |
|
859 |
using g_snd Pair_A A by (intro R.positive_integral_fst_measurable) auto |
|
860 |
then have int_g: "(\<lambda>x. \<integral>\<^isup>+ y. g y * indicator A (x, y) \<partial>M2) \<in> borel_measurable M1" |
|
861 |
by simp |
|
38656 | 862 |
|
47694 | 863 |
show "emeasure ?L A = emeasure ?R A" |
864 |
apply (subst L.emeasure_pair_measure[OF A]) |
|
865 |
apply (subst emeasure_density) |
|
866 |
using f_fst g_snd apply (simp add: split_beta') |
|
867 |
using A apply simp |
|
868 |
apply (subst positive_integral_density[OF g]) |
|
869 |
apply (simp add: indicator_eq Pair_A) |
|
870 |
apply (subst positive_integral_density[OF f]) |
|
871 |
apply (rule int_g) |
|
872 |
apply (subst R.positive_integral_fst_measurable(2)[symmetric]) |
|
873 |
using f g A Pair_A f_fst g_snd |
|
874 |
apply (auto intro!: positive_integral_cong_AE R.measurable_emeasure_Pair1 |
|
875 |
simp: positive_integral_cmult indicator_eq split_beta') |
|
876 |
apply (intro AE_I2 impI) |
|
877 |
apply (subst mult_assoc) |
|
878 |
apply (subst positive_integral_cmult) |
|
879 |
apply auto |
|
880 |
done |
|
881 |
qed simp |
|
882 |
||
883 |
lemma sigma_finite_measure_distr: |
|
884 |
assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N" |
|
885 |
shows "sigma_finite_measure M" |
|
40859 | 886 |
proof - |
47694 | 887 |
interpret sigma_finite_measure "distr M N f" by fact |
888 |
from sigma_finite_disjoint guess A . note A = this |
|
889 |
show ?thesis |
|
890 |
proof (unfold_locales, intro conjI exI allI) |
|
891 |
show "range (\<lambda>i. f -` A i \<inter> space M) \<subseteq> sets M" |
|
892 |
using A f by (auto intro!: measurable_sets) |
|
893 |
show "(\<Union>i. f -` A i \<inter> space M) = space M" |
|
894 |
using A(1) A(2)[symmetric] f by (auto simp: measurable_def Pi_def) |
|
895 |
fix i show "emeasure M (f -` A i \<inter> space M) \<noteq> \<infinity>" |
|
896 |
using f A(1,2) A(3)[of i] by (simp add: emeasure_distr subset_eq) |
|
897 |
qed |
|
38656 | 898 |
qed |
899 |
||
47694 | 900 |
lemma measurable_cong': |
901 |
assumes sets: "sets M = sets M'" "sets N = sets N'" |
|
902 |
shows "measurable M N = measurable M' N'" |
|
903 |
using sets[THEN sets_eq_imp_space_eq] sets by (simp add: measurable_def) |
|
38656 | 904 |
|
47694 | 905 |
lemma pair_measure_distr: |
906 |
assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T" |
|
907 |
assumes "sigma_finite_measure (distr M S f)" "sigma_finite_measure (distr N T g)" |
|
908 |
shows "distr M S f \<Otimes>\<^isub>M distr N T g = distr (M \<Otimes>\<^isub>M N) (S \<Otimes>\<^isub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D") |
|
909 |
proof (rule measure_eqI) |
|
910 |
show "sets ?P = sets ?D" |
|
911 |
by simp |
|
912 |
interpret S: sigma_finite_measure "distr M S f" by fact |
|
913 |
interpret T: sigma_finite_measure "distr N T g" by fact |
|
914 |
interpret ST: pair_sigma_finite "distr M S f" "distr N T g" .. |
|
915 |
interpret M: sigma_finite_measure M by (rule sigma_finite_measure_distr) fact+ |
|
916 |
interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+ |
|
917 |
interpret MN: pair_sigma_finite M N .. |
|
918 |
interpret SN: pair_sigma_finite "distr M S f" N .. |
|
919 |
have [simp]: |
|
920 |
"\<And>f g. fst \<circ> (\<lambda>(x, y). (f x, g y)) = f \<circ> fst" "\<And>f g. snd \<circ> (\<lambda>(x, y). (f x, g y)) = g \<circ> snd" |
|
921 |
by auto |
|
922 |
then have fg: "(\<lambda>(x, y). (f x, g y)) \<in> measurable (M \<Otimes>\<^isub>M N) (S \<Otimes>\<^isub>M T)" |
|
923 |
using measurable_comp[OF measurable_fst f] measurable_comp[OF measurable_snd g] |
|
924 |
by (auto simp: measurable_pair_iff) |
|
925 |
fix A assume A: "A \<in> sets ?P" |
|
926 |
then have "emeasure ?P A = (\<integral>\<^isup>+x. emeasure (distr N T g) (Pair x -` A) \<partial>distr M S f)" |
|
927 |
by (rule ST.emeasure_pair_measure_alt) |
|
928 |
also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (g -` (Pair x -` A) \<inter> space N) \<partial>distr M S f)" |
|
929 |
using g A by (simp add: sets_Pair1 emeasure_distr) |
|
930 |
also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (g -` (Pair (f x) -` A) \<inter> space N) \<partial>M)" |
|
931 |
using f g A ST.measurable_emeasure_Pair1[OF A] |
|
932 |
by (intro positive_integral_distr) (auto simp add: sets_Pair1 emeasure_distr) |
|
933 |
also have "\<dots> = (\<integral>\<^isup>+x. emeasure N (Pair x -` ((\<lambda>(x, y). (f x, g y)) -` A \<inter> space (M \<Otimes>\<^isub>M N))) \<partial>M)" |
|
934 |
by (intro positive_integral_cong arg_cong2[where f=emeasure]) (auto simp: space_pair_measure) |
|
935 |
also have "\<dots> = emeasure (M \<Otimes>\<^isub>M N) ((\<lambda>(x, y). (f x, g y)) -` A \<inter> space (M \<Otimes>\<^isub>M N))" |
|
936 |
using fg by (intro MN.emeasure_pair_measure_alt[symmetric] measurable_sets[OF _ A]) |
|
937 |
(auto cong: measurable_cong') |
|
938 |
also have "\<dots> = emeasure ?D A" |
|
939 |
using fg A by (subst emeasure_distr) auto |
|
940 |
finally show "emeasure ?P A = emeasure ?D 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:
44890
diff
changeset
|
941 |
qed |
39097 | 942 |
|
40859 | 943 |
end |