| author | wenzelm | 
| Mon, 18 Jan 2016 16:03:58 +0100 | |
| changeset 62199 | fc55a4e3f439 | 
| parent 62083 | 7582b39f51ed | 
| child 62390 | 842917225d56 | 
| 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  | 
||
| 61808 | 5  | 
section \<open>Binary product measures\<close>  | 
| 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  | 
| 
56993
 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 
hoelzl 
parents: 
54863 
diff
changeset
 | 
8  | 
imports Nonnegative_Lebesgue_Integration  | 
| 35833 | 9  | 
begin  | 
10  | 
||
| 50104 | 11  | 
lemma Pair_vimage_times[simp]: "Pair x -` (A \<times> B) = (if x \<in> A then B else {})"
 | 
| 40859 | 12  | 
by auto  | 
13  | 
||
| 50104 | 14  | 
lemma rev_Pair_vimage_times[simp]: "(\<lambda>x. (x, y)) -` (A \<times> B) = (if y \<in> B then A else {})"
 | 
| 40859 | 15  | 
by auto  | 
16  | 
||
| 56994 | 17  | 
subsection "Binary products"  | 
| 40859 | 18  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
19  | 
definition pair_measure (infixr "\<Otimes>\<^sub>M" 80) where  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
20  | 
"A \<Otimes>\<^sub>M B = measure_of (space A \<times> space B)  | 
| 47694 | 21  | 
      {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}
 | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
22  | 
(\<lambda>X. \<integral>\<^sup>+x. (\<integral>\<^sup>+y. indicator X (x,y) \<partial>B) \<partial>A)"  | 
| 40859 | 23  | 
|
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
24  | 
lemma pair_measure_closed: "{a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B} \<subseteq> Pow (space A \<times> space B)"
 | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
25  | 
using sets.space_closed[of A] sets.space_closed[of B] by auto  | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
26  | 
|
| 
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
 | 
27  | 
lemma space_pair_measure:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
28  | 
"space (A \<Otimes>\<^sub>M B) = space A \<times> space B"  | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
29  | 
unfolding pair_measure_def using pair_measure_closed[of A B]  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
30  | 
by (rule space_measure_of)  | 
| 47694 | 31  | 
|
| 59000 | 32  | 
lemma SIGMA_Collect_eq: "(SIGMA x:space M. {y\<in>space N. P x y}) = {x\<in>space (M \<Otimes>\<^sub>M N). P (fst x) (snd x)}"
 | 
33  | 
by (auto simp: space_pair_measure)  | 
|
34  | 
||
| 47694 | 35  | 
lemma sets_pair_measure:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
36  | 
  "sets (A \<Otimes>\<^sub>M B) = sigma_sets (space A \<times> space B) {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
 | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
37  | 
unfolding pair_measure_def using pair_measure_closed[of A B]  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
38  | 
by (rule sets_measure_of)  | 
| 41095 | 39  | 
|
| 58606 | 40  | 
lemma sets_pair_in_sets:  | 
41  | 
assumes N: "space A \<times> space B = space N"  | 
|
42  | 
assumes "\<And>a b. a \<in> sets A \<Longrightarrow> b \<in> sets B \<Longrightarrow> a \<times> b \<in> sets N"  | 
|
43  | 
shows "sets (A \<Otimes>\<^sub>M B) \<subseteq> sets N"  | 
|
44  | 
using assms by (auto intro!: sets.sigma_sets_subset simp: sets_pair_measure N)  | 
|
45  | 
||
| 59048 | 46  | 
lemma sets_pair_measure_cong[measurable_cong, cong]:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
47  | 
"sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^sub>M M2) = sets (M1' \<Otimes>\<^sub>M M2')"  | 
| 49776 | 48  | 
unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)  | 
49  | 
||
| 50003 | 50  | 
lemma pair_measureI[intro, simp, measurable]:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
51  | 
"x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^sub>M B)"  | 
| 47694 | 52  | 
by (auto simp: sets_pair_measure)  | 
| 41095 | 53  | 
|
| 58606 | 54  | 
lemma sets_Pair: "{x} \<in> sets M1 \<Longrightarrow> {y} \<in> sets M2 \<Longrightarrow> {(x, y)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
55  | 
  using pair_measureI[of "{x}" M1 "{y}" M2] by simp
 | 
|
56  | 
||
| 47694 | 57  | 
lemma measurable_pair_measureI:  | 
58  | 
assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2"  | 
|
59  | 
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"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
60  | 
shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"  | 
| 47694 | 61  | 
unfolding pair_measure_def using 1 2  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
62  | 
by (intro measurable_measure_of) (auto dest: sets.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
 | 
63  | 
|
| 50003 | 64  | 
lemma measurable_split_replace[measurable (raw)]:  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61169 
diff
changeset
 | 
65  | 
"(\<lambda>x. f x (fst (g x)) (snd (g x))) \<in> measurable M N \<Longrightarrow> (\<lambda>x. case_prod (f x) (g x)) \<in> measurable M N"  | 
| 50003 | 66  | 
unfolding split_beta' .  | 
67  | 
||
68  | 
lemma measurable_Pair[measurable (raw)]:  | 
|
| 49776 | 69  | 
assumes f: "f \<in> measurable M M1" and g: "g \<in> measurable M M2"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
70  | 
shows "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"  | 
| 49776 | 71  | 
proof (rule measurable_pair_measureI)  | 
72  | 
show "(\<lambda>x. (f x, g x)) \<in> space M \<rightarrow> space M1 \<times> space M2"  | 
|
73  | 
using f g by (auto simp: measurable_def)  | 
|
74  | 
fix A B assume *: "A \<in> sets M1" "B \<in> sets M2"  | 
|
75  | 
have "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"  | 
|
76  | 
by auto  | 
|
77  | 
also have "\<dots> \<in> sets M"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
78  | 
by (rule sets.Int) (auto intro!: measurable_sets * f g)  | 
| 49776 | 79  | 
finally show "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M \<in> sets M" .  | 
80  | 
qed  | 
|
81  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
82  | 
lemma measurable_fst[intro!, simp, measurable]: "fst \<in> measurable (M1 \<Otimes>\<^sub>M M2) M1"  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
83  | 
by (auto simp: fst_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times  | 
| 
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
84  | 
measurable_def)  | 
| 40859 | 85  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
86  | 
lemma measurable_snd[intro!, simp, measurable]: "snd \<in> measurable (M1 \<Otimes>\<^sub>M M2) M2"  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
87  | 
by (auto simp: snd_vimage_eq_Times space_pair_measure sets.sets_into_space times_Int_times  | 
| 
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
88  | 
measurable_def)  | 
| 47694 | 89  | 
|
| 
59353
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
90  | 
lemma measurable_Pair_compose_split[measurable_dest]:  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61169 
diff
changeset
 | 
91  | 
assumes f: "case_prod f \<in> measurable (M1 \<Otimes>\<^sub>M M2) N"  | 
| 
59353
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
92  | 
assumes g: "g \<in> measurable M M1" and h: "h \<in> measurable M M2"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
93  | 
shows "(\<lambda>x. f (g x) (h x)) \<in> measurable M N"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
94  | 
using measurable_compose[OF measurable_Pair f, OF g h] by simp  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
95  | 
|
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
96  | 
lemma measurable_Pair1_compose[measurable_dest]:  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
97  | 
assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
98  | 
assumes [measurable]: "h \<in> measurable N M"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
99  | 
shows "(\<lambda>x. f (h x)) \<in> measurable N M1"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
100  | 
using measurable_compose[OF f measurable_fst] by simp  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
101  | 
|
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
102  | 
lemma measurable_Pair2_compose[measurable_dest]:  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
103  | 
assumes f: "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
104  | 
assumes [measurable]: "h \<in> measurable N M"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
105  | 
shows "(\<lambda>x. g (h x)) \<in> measurable N M2"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
106  | 
using measurable_compose[OF f measurable_snd] by simp  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
107  | 
|
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
108  | 
lemma measurable_pair:  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
109  | 
assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<in> measurable M M2"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
110  | 
shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
111  | 
using measurable_Pair[OF assms] by simp  | 
| 
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
112  | 
|
| 50003 | 113  | 
lemma  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
114  | 
assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)"  | 
| 50003 | 115  | 
shows measurable_fst': "(\<lambda>x. fst (f x)) \<in> measurable M N"  | 
116  | 
and measurable_snd': "(\<lambda>x. snd (f x)) \<in> measurable M P"  | 
|
117  | 
by simp_all  | 
|
| 40859 | 118  | 
|
| 50003 | 119  | 
lemma  | 
120  | 
assumes f[measurable]: "f \<in> measurable M N"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
121  | 
shows measurable_fst'': "(\<lambda>x. f (fst x)) \<in> measurable (M \<Otimes>\<^sub>M P) N"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
122  | 
and measurable_snd'': "(\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^sub>M M) N"  | 
| 50003 | 123  | 
by simp_all  | 
| 47694 | 124  | 
|
| 58606 | 125  | 
lemma sets_pair_eq_sets_fst_snd:  | 
126  | 
  "sets (A \<Otimes>\<^sub>M B) = sets (Sup_sigma {vimage_algebra (space A \<times> space B) fst A, vimage_algebra (space A \<times> space B) snd B})"
 | 
|
127  | 
    (is "?P = sets (Sup_sigma {?fst, ?snd})")
 | 
|
128  | 
proof -  | 
|
129  | 
  { fix a b assume ab: "a \<in> sets A" "b \<in> sets B"
 | 
|
130  | 
then have "a \<times> b = (fst -` a \<inter> (space A \<times> space B)) \<inter> (snd -` b \<inter> (space A \<times> space B))"  | 
|
131  | 
by (auto dest: sets.sets_into_space)  | 
|
132  | 
    also have "\<dots> \<in> sets (Sup_sigma {?fst, ?snd})"
 | 
|
133  | 
using ab by (auto intro: in_Sup_sigma in_vimage_algebra)  | 
|
134  | 
    finally have "a \<times> b \<in> sets (Sup_sigma {?fst, ?snd})" . }
 | 
|
135  | 
moreover have "sets ?fst \<subseteq> sets (A \<Otimes>\<^sub>M B)"  | 
|
136  | 
by (rule sets_image_in_sets) (auto simp: space_pair_measure[symmetric])  | 
|
137  | 
moreover have "sets ?snd \<subseteq> sets (A \<Otimes>\<^sub>M B)"  | 
|
138  | 
by (rule sets_image_in_sets) (auto simp: space_pair_measure)  | 
|
139  | 
ultimately show ?thesis  | 
|
140  | 
by (intro antisym[of "sets A" for A] sets_Sup_in_sets sets_pair_in_sets )  | 
|
141  | 
(auto simp add: space_Sup_sigma space_pair_measure)  | 
|
142  | 
qed  | 
|
143  | 
||
| 47694 | 144  | 
lemma measurable_pair_iff:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
145  | 
"f \<in> measurable M (M1 \<Otimes>\<^sub>M M2) \<longleftrightarrow> (fst \<circ> f) \<in> measurable M M1 \<and> (snd \<circ> f) \<in> measurable M M2"  | 
| 50003 | 146  | 
by (auto intro: measurable_pair[of f M M1 M2])  | 
| 40859 | 147  | 
|
| 49776 | 148  | 
lemma measurable_split_conv:  | 
149  | 
"(\<lambda>(x, y). f x y) \<in> measurable A B \<longleftrightarrow> (\<lambda>x. f (fst x) (snd x)) \<in> measurable A B"  | 
|
150  | 
by (intro arg_cong2[where f="op \<in>"]) auto  | 
|
| 40859 | 151  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
152  | 
lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)"  | 
| 49776 | 153  | 
by (auto intro!: measurable_Pair simp: measurable_split_conv)  | 
| 47694 | 154  | 
|
155  | 
lemma measurable_pair_swap:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
156  | 
assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" shows "(\<lambda>(x,y). f (y, x)) \<in> measurable (M2 \<Otimes>\<^sub>M M1) M"  | 
| 49776 | 157  | 
using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def)  | 
| 40859 | 158  | 
|
| 47694 | 159  | 
lemma measurable_pair_swap_iff:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
160  | 
"f \<in> measurable (M2 \<Otimes>\<^sub>M M1) M \<longleftrightarrow> (\<lambda>(x,y). f (y,x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) M"  | 
| 50003 | 161  | 
by (auto dest: measurable_pair_swap)  | 
| 49776 | 162  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
163  | 
lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)"  | 
| 50003 | 164  | 
by simp  | 
| 40859 | 165  | 
|
| 50003 | 166  | 
lemma sets_Pair1[measurable (raw)]:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
167  | 
assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "Pair x -` A \<in> sets M2"  | 
| 40859 | 168  | 
proof -  | 
| 47694 | 169  | 
  have "Pair x -` A = (if x \<in> space M1 then Pair x -` A \<inter> space M2 else {})"
 | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
170  | 
using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)  | 
| 47694 | 171  | 
also have "\<dots> \<in> sets M2"  | 
172  | 
using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: split_if_asm)  | 
|
173  | 
finally show ?thesis .  | 
|
| 40859 | 174  | 
qed  | 
175  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
176  | 
lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)"  | 
| 49776 | 177  | 
by (auto intro!: measurable_Pair)  | 
| 40859 | 178  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
179  | 
lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"  | 
| 47694 | 180  | 
proof -  | 
181  | 
  have "(\<lambda>x. (x, y)) -` A = (if y \<in> space M2 then (\<lambda>x. (x, y)) -` A \<inter> space M1 else {})"
 | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
182  | 
using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)  | 
| 47694 | 183  | 
also have "\<dots> \<in> sets M1"  | 
184  | 
using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: split_if_asm)  | 
|
185  | 
finally show ?thesis .  | 
|
| 40859 | 186  | 
qed  | 
187  | 
||
| 47694 | 188  | 
lemma measurable_Pair2:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
189  | 
assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and x: "x \<in> space M1"  | 
| 47694 | 190  | 
shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"  | 
191  | 
using measurable_comp[OF measurable_Pair1' f, OF x]  | 
|
192  | 
by (simp add: comp_def)  | 
|
193  | 
||
194  | 
lemma measurable_Pair1:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
195  | 
assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and y: "y \<in> space M2"  | 
| 40859 | 196  | 
shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"  | 
| 47694 | 197  | 
using measurable_comp[OF measurable_Pair2' f, OF y]  | 
198  | 
by (simp add: comp_def)  | 
|
| 40859 | 199  | 
|
| 47694 | 200  | 
lemma Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
 | 
| 40859 | 201  | 
unfolding Int_stable_def  | 
| 47694 | 202  | 
by safe (auto simp add: times_Int_times)  | 
| 40859 | 203  | 
|
| 49776 | 204  | 
lemma (in finite_measure) finite_measure_cut_measurable:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
205  | 
assumes [measurable]: "Q \<in> sets (N \<Otimes>\<^sub>M M)"  | 
| 49776 | 206  | 
shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N"  | 
| 40859 | 207  | 
(is "?s Q \<in> _")  | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
208  | 
using Int_stable_pair_measure_generator pair_measure_closed assms  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
209  | 
unfolding sets_pair_measure  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
210  | 
proof (induct rule: sigma_sets_induct_disjoint)  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
211  | 
case (compl A)  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
212  | 
with sets.sets_into_space have "\<And>x. emeasure M (Pair x -` ((space N \<times> space M) - A)) =  | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
213  | 
(if x \<in> space N then emeasure M (space M) - ?s A x else 0)"  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
214  | 
unfolding sets_pair_measure[symmetric]  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
215  | 
by (auto intro!: emeasure_compl simp: vimage_Diff sets_Pair1)  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
216  | 
with compl sets.top show ?case  | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
217  | 
by (auto intro!: measurable_If simp: space_pair_measure)  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
218  | 
next  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
219  | 
case (union F)  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
220  | 
then have "\<And>x. emeasure M (Pair x -` (\<Union>i. F i)) = (\<Sum>i. ?s (F i) x)"  | 
| 60727 | 221  | 
by (simp add: suminf_emeasure disjoint_family_on_vimageI subset_eq vimage_UN sets_pair_measure[symmetric])  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
222  | 
with union show ?case  | 
| 50003 | 223  | 
unfolding sets_pair_measure[symmetric] by simp  | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
224  | 
qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If)  | 
| 49776 | 225  | 
|
226  | 
lemma (in sigma_finite_measure) measurable_emeasure_Pair:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
227  | 
assumes Q: "Q \<in> sets (N \<Otimes>\<^sub>M M)" shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N" (is "?s Q \<in> _")  | 
| 49776 | 228  | 
proof -  | 
229  | 
from sigma_finite_disjoint guess F . note F = this  | 
|
230  | 
then have F_sets: "\<And>i. F i \<in> sets M" by auto  | 
|
231  | 
let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q"  | 
|
232  | 
  { fix i
 | 
|
233  | 
have [simp]: "space N \<times> F i \<inter> space N \<times> space M = space N \<times> F i"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
234  | 
using F sets.sets_into_space by auto  | 
| 49776 | 235  | 
let ?R = "density M (indicator (F i))"  | 
236  | 
have "finite_measure ?R"  | 
|
237  | 
using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq)  | 
|
238  | 
then have "(\<lambda>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))) \<in> borel_measurable N"  | 
|
239  | 
by (rule finite_measure.finite_measure_cut_measurable) (auto intro: Q)  | 
|
240  | 
moreover have "\<And>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))  | 
|
241  | 
= emeasure M (F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q))"  | 
|
242  | 
using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1)  | 
|
243  | 
moreover have "\<And>x. F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q) = ?C x i"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
244  | 
using sets.sets_into_space[OF Q] by (auto simp: space_pair_measure)  | 
| 49776 | 245  | 
ultimately have "(\<lambda>x. emeasure M (?C x i)) \<in> borel_measurable N"  | 
246  | 
by simp }  | 
|
247  | 
moreover  | 
|
248  | 
  { fix x
 | 
|
249  | 
have "(\<Sum>i. emeasure M (?C x i)) = emeasure M (\<Union>i. ?C x i)"  | 
|
250  | 
proof (intro suminf_emeasure)  | 
|
251  | 
show "range (?C x) \<subseteq> sets M"  | 
|
| 61808 | 252  | 
using F \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> by (auto intro!: sets_Pair1)  | 
| 49776 | 253  | 
have "disjoint_family F" using F by auto  | 
254  | 
show "disjoint_family (?C x)"  | 
|
| 61808 | 255  | 
by (rule disjoint_family_on_bisimulation[OF \<open>disjoint_family F\<close>]) auto  | 
| 49776 | 256  | 
qed  | 
257  | 
also have "(\<Union>i. ?C x i) = Pair x -` Q"  | 
|
| 61808 | 258  | 
using F sets.sets_into_space[OF \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close>]  | 
| 49776 | 259  | 
by (auto simp: space_pair_measure)  | 
260  | 
finally have "emeasure M (Pair x -` Q) = (\<Sum>i. emeasure M (?C x i))"  | 
|
261  | 
by simp }  | 
|
| 61808 | 262  | 
ultimately show ?thesis using \<open>Q \<in> sets (N \<Otimes>\<^sub>M M)\<close> F_sets  | 
| 49776 | 263  | 
by auto  | 
264  | 
qed  | 
|
265  | 
||
| 50003 | 266  | 
lemma (in sigma_finite_measure) measurable_emeasure[measurable (raw)]:  | 
267  | 
assumes space: "\<And>x. x \<in> space N \<Longrightarrow> A x \<subseteq> space M"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
268  | 
  assumes A: "{x\<in>space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M)"
 | 
| 50003 | 269  | 
shows "(\<lambda>x. emeasure M (A x)) \<in> borel_measurable N"  | 
270  | 
proof -  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
271  | 
  from space have "\<And>x. x \<in> space N \<Longrightarrow> Pair x -` {x \<in> space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} = A x"
 | 
| 50003 | 272  | 
by (auto simp: space_pair_measure)  | 
273  | 
with measurable_emeasure_Pair[OF A] show ?thesis  | 
|
274  | 
by (auto cong: measurable_cong)  | 
|
275  | 
qed  | 
|
276  | 
||
| 49776 | 277  | 
lemma (in sigma_finite_measure) emeasure_pair_measure:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
278  | 
assumes "X \<in> sets (N \<Otimes>\<^sub>M M)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
279  | 
shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator X (x, y) \<partial>M \<partial>N)" (is "_ = ?\<mu> X")  | 
| 49776 | 280  | 
proof (rule emeasure_measure_of[OF pair_measure_def])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
281  | 
show "positive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"  | 
| 56996 | 282  | 
by (auto simp: positive_def nn_integral_nonneg)  | 
| 49776 | 283  | 
have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y"  | 
284  | 
by (auto simp: indicator_def)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
285  | 
show "countably_additive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"  | 
| 49776 | 286  | 
proof (rule countably_additiveI)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
287  | 
    fix F :: "nat \<Rightarrow> ('b \<times> 'a) set" assume F: "range F \<subseteq> sets (N \<Otimes>\<^sub>M M)" "disjoint_family F"
 | 
| 
59353
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
288  | 
from F have *: "\<And>i. F i \<in> sets (N \<Otimes>\<^sub>M M)" by auto  | 
| 49776 | 289  | 
moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"  | 
290  | 
by (intro disjoint_family_on_bisimulation[OF F(2)]) auto  | 
|
291  | 
moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M"  | 
|
292  | 
using F by (auto simp: sets_Pair1)  | 
|
293  | 
ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)"  | 
|
| 
59353
 
f0707dc3d9aa
measurability prover: removed app splitting, replaced by more powerful destruction rules
 
hoelzl 
parents: 
59088 
diff
changeset
 | 
294  | 
by (auto simp add: nn_integral_suminf[symmetric] vimage_UN suminf_emeasure emeasure_nonneg  | 
| 56996 | 295  | 
intro!: nn_integral_cong nn_integral_indicator[symmetric])  | 
| 49776 | 296  | 
qed  | 
297  | 
  show "{a \<times> b |a b. a \<in> sets N \<and> b \<in> sets M} \<subseteq> Pow (space N \<times> space M)"
 | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
298  | 
using sets.space_closed[of N] sets.space_closed[of M] by auto  | 
| 49776 | 299  | 
qed fact  | 
300  | 
||
301  | 
lemma (in sigma_finite_measure) emeasure_pair_measure_alt:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
302  | 
assumes X: "X \<in> sets (N \<Otimes>\<^sub>M M)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
303  | 
shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+x. emeasure M (Pair x -` X) \<partial>N)"  | 
| 49776 | 304  | 
proof -  | 
305  | 
have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y"  | 
|
306  | 
by (auto simp: indicator_def)  | 
|
307  | 
show ?thesis  | 
|
| 56996 | 308  | 
using X by (auto intro!: nn_integral_cong simp: emeasure_pair_measure sets_Pair1)  | 
| 49776 | 309  | 
qed  | 
310  | 
||
311  | 
lemma (in sigma_finite_measure) emeasure_pair_measure_Times:  | 
|
312  | 
assumes A: "A \<in> sets N" and B: "B \<in> sets M"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
313  | 
shows "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = emeasure N A * emeasure M B"  | 
| 49776 | 314  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
315  | 
have "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = (\<integral>\<^sup>+x. emeasure M B * indicator A x \<partial>N)"  | 
| 56996 | 316  | 
using A B by (auto intro!: nn_integral_cong simp: emeasure_pair_measure_alt)  | 
| 49776 | 317  | 
also have "\<dots> = emeasure M B * emeasure N A"  | 
| 56996 | 318  | 
using A by (simp add: emeasure_nonneg nn_integral_cmult_indicator)  | 
| 49776 | 319  | 
finally show ?thesis  | 
320  | 
by (simp add: ac_simps)  | 
|
| 40859 | 321  | 
qed  | 
322  | 
||
| 61808 | 323  | 
subsection \<open>Binary products of $\sigma$-finite emeasure spaces\<close>  | 
| 40859 | 324  | 
|
| 
61565
 
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
 
ballarin 
parents: 
61424 
diff
changeset
 | 
325  | 
locale pair_sigma_finite = M1?: sigma_finite_measure M1 + M2?: sigma_finite_measure M2  | 
| 47694 | 326  | 
for M1 :: "'a measure" and M2 :: "'b measure"  | 
| 40859 | 327  | 
|
| 47694 | 328  | 
lemma (in pair_sigma_finite) measurable_emeasure_Pair1:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
329  | 
"Q \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow> (\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"  | 
| 49776 | 330  | 
using M2.measurable_emeasure_Pair .  | 
| 40859 | 331  | 
|
| 47694 | 332  | 
lemma (in pair_sigma_finite) measurable_emeasure_Pair2:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
333  | 
assumes Q: "Q \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>y. emeasure M1 ((\<lambda>x. (x, y)) -` Q)) \<in> borel_measurable M2"  | 
| 40859 | 334  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
335  | 
have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"  | 
| 47694 | 336  | 
using Q measurable_pair_swap' by (auto intro: measurable_sets)  | 
| 49776 | 337  | 
note M1.measurable_emeasure_Pair[OF this]  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
338  | 
moreover have "\<And>y. Pair y -` ((\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1)) = (\<lambda>x. (x, y)) -` Q"  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
339  | 
using Q[THEN sets.sets_into_space] by (auto simp: space_pair_measure)  | 
| 47694 | 340  | 
ultimately show ?thesis by simp  | 
| 
39088
 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 
hoelzl 
parents: 
39082 
diff
changeset
 | 
341  | 
qed  | 
| 
 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 
hoelzl 
parents: 
39082 
diff
changeset
 | 
342  | 
|
| 
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
 | 
343  | 
lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:  | 
| 47694 | 344  | 
  defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}"
 | 
345  | 
  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>
 | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
346  | 
(\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)"  | 
| 40859 | 347  | 
proof -  | 
| 47694 | 348  | 
from M1.sigma_finite_incseq guess F1 . note F1 = this  | 
349  | 
from M2.sigma_finite_incseq guess F2 . note F2 = this  | 
|
350  | 
from F1 F2 have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto  | 
|
| 40859 | 351  | 
let ?F = "\<lambda>i. F1 i \<times> F2 i"  | 
| 47694 | 352  | 
show ?thesis  | 
| 40859 | 353  | 
proof (intro exI[of _ ?F] conjI allI)  | 
| 47694 | 354  | 
show "range ?F \<subseteq> E" using F1 F2 by (auto simp: E_def) (metis range_subsetD)  | 
| 40859 | 355  | 
next  | 
356  | 
have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"  | 
|
357  | 
proof (intro subsetI)  | 
|
358  | 
fix x assume "x \<in> space M1 \<times> space M2"  | 
|
359  | 
then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j"  | 
|
360  | 
by (auto simp: space)  | 
|
361  | 
then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"  | 
|
| 61808 | 362  | 
using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_def  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
363  | 
by (force split: split_max)+  | 
| 40859 | 364  | 
then have "(fst x, snd x) \<in> F1 (max i j) \<times> F2 (max i j)"  | 
| 
54863
 
82acc20ded73
prefer more canonical names for lemmas on min/max
 
haftmann 
parents: 
53374 
diff
changeset
 | 
365  | 
by (intro SigmaI) (auto simp add: max.commute)  | 
| 40859 | 366  | 
then show "x \<in> (\<Union>i. ?F i)" by auto  | 
367  | 
qed  | 
|
| 47694 | 368  | 
then show "(\<Union>i. ?F i) = space M1 \<times> space M2"  | 
369  | 
using space by (auto simp: space)  | 
|
| 40859 | 370  | 
next  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
371  | 
fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)"  | 
| 61808 | 372  | 
using \<open>incseq F1\<close> \<open>incseq F2\<close> unfolding incseq_Suc_iff by auto  | 
| 40859 | 373  | 
next  | 
374  | 
fix i  | 
|
375  | 
from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto  | 
|
| 47694 | 376  | 
with F1 F2 emeasure_nonneg[of M1 "F1 i"] emeasure_nonneg[of M2 "F2 i"]  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
377  | 
show "emeasure (M1 \<Otimes>\<^sub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"  | 
| 47694 | 378  | 
by (auto simp add: emeasure_pair_measure_Times)  | 
379  | 
qed  | 
|
380  | 
qed  | 
|
381  | 
||
| 
61565
 
352c73a689da
Qualifiers in locale expressions default to mandatory regardless of the command.
 
ballarin 
parents: 
61424 
diff
changeset
 | 
382  | 
sublocale pair_sigma_finite \<subseteq> P?: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"  | 
| 47694 | 383  | 
proof  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
384  | 
from M1.sigma_finite_countable guess F1 ..  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
385  | 
moreover from M2.sigma_finite_countable guess F2 ..  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
386  | 
ultimately show  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
387  | 
"\<exists>A. countable A \<and> A \<subseteq> sets (M1 \<Otimes>\<^sub>M M2) \<and> \<Union>A = space (M1 \<Otimes>\<^sub>M M2) \<and> (\<forall>a\<in>A. emeasure (M1 \<Otimes>\<^sub>M M2) a \<noteq> \<infinity>)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
388  | 
by (intro exI[of _ "(\<lambda>(a, b). a \<times> b) ` (F1 \<times> F2)"] conjI)  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
389  | 
(auto simp: M2.emeasure_pair_measure_Times space_pair_measure set_eq_iff subset_eq  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
390  | 
dest: sets.sets_into_space)  | 
| 40859 | 391  | 
qed  | 
392  | 
||
| 47694 | 393  | 
lemma sigma_finite_pair_measure:  | 
394  | 
assumes A: "sigma_finite_measure A" and B: "sigma_finite_measure B"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
395  | 
shows "sigma_finite_measure (A \<Otimes>\<^sub>M B)"  | 
| 47694 | 396  | 
proof -  | 
397  | 
interpret A: sigma_finite_measure A by fact  | 
|
398  | 
interpret B: sigma_finite_measure B by fact  | 
|
399  | 
interpret AB: pair_sigma_finite A B ..  | 
|
400  | 
show ?thesis ..  | 
|
| 40859 | 401  | 
qed  | 
| 
39088
 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 
hoelzl 
parents: 
39082 
diff
changeset
 | 
402  | 
|
| 47694 | 403  | 
lemma sets_pair_swap:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
404  | 
assumes "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
405  | 
shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"  | 
| 47694 | 406  | 
using measurable_pair_swap' assms by (rule measurable_sets)  | 
| 41661 | 407  | 
|
| 47694 | 408  | 
lemma (in pair_sigma_finite) distr_pair_swap:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
409  | 
"M1 \<Otimes>\<^sub>M M2 = distr (M2 \<Otimes>\<^sub>M M1) (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x, y). (y, x))" (is "?P = ?D")  | 
| 40859 | 410  | 
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
 | 
411  | 
  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
 | 
| 47694 | 412  | 
  let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
 | 
413  | 
show ?thesis  | 
|
414  | 
proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])  | 
|
415  | 
show "?E \<subseteq> Pow (space ?P)"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
416  | 
using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure)  | 
| 47694 | 417  | 
show "sets ?P = sigma_sets (space ?P) ?E"  | 
418  | 
by (simp add: sets_pair_measure space_pair_measure)  | 
|
419  | 
then show "sets ?D = sigma_sets (space ?P) ?E"  | 
|
420  | 
by simp  | 
|
421  | 
next  | 
|
| 
49784
 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 
hoelzl 
parents: 
49776 
diff
changeset
 | 
422  | 
show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"  | 
| 47694 | 423  | 
using F by (auto simp: space_pair_measure)  | 
424  | 
next  | 
|
425  | 
fix X assume "X \<in> ?E"  | 
|
426  | 
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  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
427  | 
have "(\<lambda>(y, x). (x, y)) -` X \<inter> space (M2 \<Otimes>\<^sub>M M1) = B \<times> A"  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
428  | 
using sets.sets_into_space[OF A] sets.sets_into_space[OF B] by (auto simp: space_pair_measure)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
429  | 
with A B show "emeasure (M1 \<Otimes>\<^sub>M M2) X = emeasure ?D X"  | 
| 49776 | 430  | 
by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr  | 
| 47694 | 431  | 
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
 | 
432  | 
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
 | 
433  | 
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
 | 
434  | 
|
| 47694 | 435  | 
lemma (in pair_sigma_finite) emeasure_pair_measure_alt2:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
436  | 
assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
437  | 
shows "emeasure (M1 \<Otimes>\<^sub>M M2) A = (\<integral>\<^sup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"  | 
| 47694 | 438  | 
(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
 | 
439  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
440  | 
have [simp]: "\<And>y. (Pair y -` ((\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1))) = (\<lambda>x. (x, y)) -` A"  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
441  | 
using sets.sets_into_space[OF A] by (auto simp: space_pair_measure)  | 
| 47694 | 442  | 
show ?thesis using A  | 
443  | 
by (subst distr_pair_swap)  | 
|
444  | 
(simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap']  | 
|
| 49776 | 445  | 
M1.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A])  | 
446  | 
qed  | 
|
447  | 
||
448  | 
lemma (in pair_sigma_finite) AE_pair:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
449  | 
assumes "AE x in (M1 \<Otimes>\<^sub>M M2). Q x"  | 
| 49776 | 450  | 
shows "AE x in M1. (AE y in M2. Q (x, y))"  | 
451  | 
proof -  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
452  | 
  obtain N where N: "N \<in> sets (M1 \<Otimes>\<^sub>M M2)" "emeasure (M1 \<Otimes>\<^sub>M M2) N = 0" "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> Q x} \<subseteq> N"
 | 
| 49776 | 453  | 
using assms unfolding eventually_ae_filter by auto  | 
454  | 
show ?thesis  | 
|
455  | 
proof (rule AE_I)  | 
|
| 61808 | 456  | 
from N measurable_emeasure_Pair1[OF \<open>N \<in> sets (M1 \<Otimes>\<^sub>M M2)\<close>]  | 
| 49776 | 457  | 
    show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
 | 
| 56996 | 458  | 
by (auto simp: M2.emeasure_pair_measure_alt nn_integral_0_iff emeasure_nonneg)  | 
| 49776 | 459  | 
    show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
 | 
460  | 
by (intro borel_measurable_ereal_neq_const measurable_emeasure_Pair1 N)  | 
|
461  | 
    { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0"
 | 
|
462  | 
have "AE y in M2. Q (x, y)"  | 
|
463  | 
proof (rule AE_I)  | 
|
464  | 
show "emeasure M2 (Pair x -` N) = 0" by fact  | 
|
465  | 
show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1)  | 
|
466  | 
        show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
 | 
|
| 61808 | 467  | 
using N \<open>x \<in> space M1\<close> unfolding space_pair_measure by auto  | 
| 49776 | 468  | 
qed }  | 
469  | 
    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}"
 | 
|
470  | 
by auto  | 
|
471  | 
qed  | 
|
472  | 
qed  | 
|
473  | 
||
474  | 
lemma (in pair_sigma_finite) AE_pair_measure:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
475  | 
  assumes "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
| 49776 | 476  | 
assumes ae: "AE x in M1. AE y in M2. P (x, y)"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
477  | 
shows "AE x in M1 \<Otimes>\<^sub>M M2. P x"  | 
| 49776 | 478  | 
proof (subst AE_iff_measurable[OF _ refl])  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
479  | 
  show "{x\<in>space (M1 \<Otimes>\<^sub>M M2). \<not> P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
480  | 
by (rule sets.sets_Collect) fact  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
481  | 
  then have "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} =
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
482  | 
      (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} (x, y) \<partial>M2 \<partial>M1)"
 | 
| 49776 | 483  | 
by (simp add: M2.emeasure_pair_measure)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
484  | 
also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. 0 \<partial>M2 \<partial>M1)"  | 
| 49776 | 485  | 
using ae  | 
| 56996 | 486  | 
apply (safe intro!: nn_integral_cong_AE)  | 
| 49776 | 487  | 
apply (intro AE_I2)  | 
| 56996 | 488  | 
apply (safe intro!: nn_integral_cong_AE)  | 
| 49776 | 489  | 
apply auto  | 
490  | 
done  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
491  | 
  finally show "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} = 0" by simp
 | 
| 49776 | 492  | 
qed  | 
493  | 
||
494  | 
lemma (in pair_sigma_finite) AE_pair_iff:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
495  | 
  "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow>
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
496  | 
(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE x in (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x))"  | 
| 49776 | 497  | 
using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto  | 
498  | 
||
499  | 
lemma (in pair_sigma_finite) AE_commute:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
500  | 
  assumes P: "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
| 49776 | 501  | 
shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)"  | 
502  | 
proof -  | 
|
503  | 
interpret Q: pair_sigma_finite M2 M1 ..  | 
|
504  | 
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"  | 
|
505  | 
by auto  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
506  | 
  have "{x \<in> space (M2 \<Otimes>\<^sub>M M1). P (snd x) (fst x)} =
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
507  | 
    (\<lambda>(x, y). (y, x)) -` {x \<in> space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<inter> space (M2 \<Otimes>\<^sub>M M1)"
 | 
| 49776 | 508  | 
by (auto simp: space_pair_measure)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
509  | 
also have "\<dots> \<in> sets (M2 \<Otimes>\<^sub>M M1)"  | 
| 49776 | 510  | 
by (intro sets_pair_swap P)  | 
511  | 
finally show ?thesis  | 
|
512  | 
apply (subst AE_pair_iff[OF P])  | 
|
513  | 
apply (subst distr_pair_swap)  | 
|
514  | 
apply (subst AE_distr_iff[OF measurable_pair_swap' P])  | 
|
515  | 
apply (subst Q.AE_pair_iff)  | 
|
516  | 
apply simp_all  | 
|
517  | 
done  | 
|
| 40859 | 518  | 
qed  | 
519  | 
||
| 56994 | 520  | 
subsection "Fubinis theorem"  | 
| 40859 | 521  | 
|
| 49800 | 522  | 
lemma measurable_compose_Pair1:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
523  | 
"x \<in> space M1 \<Longrightarrow> g \<in> measurable (M1 \<Otimes>\<^sub>M M2) L \<Longrightarrow> (\<lambda>y. g (x, y)) \<in> measurable M2 L"  | 
| 50003 | 524  | 
by simp  | 
| 49800 | 525  | 
|
| 56996 | 526  | 
lemma (in sigma_finite_measure) borel_measurable_nn_integral_fst':  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
527  | 
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" "\<And>x. 0 \<le> f x"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
528  | 
shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"  | 
| 49800 | 529  | 
using f proof induct  | 
530  | 
case (cong u v)  | 
|
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
531  | 
then have "\<And>w x. w \<in> space M1 \<Longrightarrow> x \<in> space M \<Longrightarrow> u (w, x) = v (w, x)"  | 
| 49800 | 532  | 
by (auto simp: space_pair_measure)  | 
533  | 
show ?case  | 
|
534  | 
apply (subst measurable_cong)  | 
|
| 56996 | 535  | 
apply (rule nn_integral_cong)  | 
| 49800 | 536  | 
apply fact+  | 
537  | 
done  | 
|
538  | 
next  | 
|
539  | 
case (set Q)  | 
|
540  | 
have [simp]: "\<And>x y. indicator Q (x, y) = indicator (Pair x -` Q) y"  | 
|
541  | 
by (auto simp: indicator_def)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
542  | 
have "\<And>x. x \<in> space M1 \<Longrightarrow> emeasure M (Pair x -` Q) = \<integral>\<^sup>+ y. indicator Q (x, y) \<partial>M"  | 
| 49800 | 543  | 
by (simp add: sets_Pair1[OF set])  | 
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
544  | 
from this measurable_emeasure_Pair[OF set] show ?case  | 
| 49800 | 545  | 
by (rule measurable_cong[THEN iffD1])  | 
| 56996 | 546  | 
qed (simp_all add: nn_integral_add nn_integral_cmult measurable_compose_Pair1  | 
547  | 
nn_integral_monotone_convergence_SUP incseq_def le_fun_def  | 
|
| 49800 | 548  | 
cong: measurable_cong)  | 
549  | 
||
| 56996 | 550  | 
lemma (in sigma_finite_measure) nn_integral_fst':  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
551  | 
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)" "\<And>x. 0 \<le> f x"  | 
| 56996 | 552  | 
shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f" (is "?I f = _")  | 
| 49800 | 553  | 
using f proof induct  | 
554  | 
case (cong u v)  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
555  | 
then have "?I u = ?I v"  | 
| 56996 | 556  | 
by (intro nn_integral_cong) (auto simp: space_pair_measure)  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
557  | 
with cong show ?case  | 
| 56996 | 558  | 
by (simp cong: nn_integral_cong)  | 
559  | 
qed (simp_all add: emeasure_pair_measure nn_integral_cmult nn_integral_add  | 
|
560  | 
nn_integral_monotone_convergence_SUP  | 
|
561  | 
measurable_compose_Pair1 nn_integral_nonneg  | 
|
562  | 
borel_measurable_nn_integral_fst' nn_integral_mono incseq_def le_fun_def  | 
|
563  | 
cong: nn_integral_cong)  | 
|
| 40859 | 564  | 
|
| 56996 | 565  | 
lemma (in sigma_finite_measure) nn_integral_fst:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
566  | 
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M)"  | 
| 56996 | 567  | 
shows "(\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M) \<partial>M1) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M) f"  | 
| 49800 | 568  | 
using f  | 
| 56996 | 569  | 
borel_measurable_nn_integral_fst'[of "\<lambda>x. max 0 (f x)"]  | 
570  | 
nn_integral_fst'[of "\<lambda>x. max 0 (f x)"]  | 
|
571  | 
unfolding nn_integral_max_0 by auto  | 
|
| 40859 | 572  | 
|
| 56996 | 573  | 
lemma (in sigma_finite_measure) borel_measurable_nn_integral[measurable (raw)]:  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61169 
diff
changeset
 | 
574  | 
"case_prod f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^sup>+ y. f x y \<partial>M) \<in> borel_measurable N"  | 
| 
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61169 
diff
changeset
 | 
575  | 
using borel_measurable_nn_integral_fst'[of "\<lambda>x. max 0 (case_prod f x)" N]  | 
| 56996 | 576  | 
by (simp add: nn_integral_max_0)  | 
| 50003 | 577  | 
|
| 56996 | 578  | 
lemma (in pair_sigma_finite) nn_integral_snd:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
579  | 
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"  | 
| 56996 | 580  | 
shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f"  | 
| 41661 | 581  | 
proof -  | 
| 47694 | 582  | 
note measurable_pair_swap[OF f]  | 
| 56996 | 583  | 
from M1.nn_integral_fst[OF this]  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
584  | 
have "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1))"  | 
| 40859 | 585  | 
by simp  | 
| 56996 | 586  | 
also have "(\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>N (M1 \<Otimes>\<^sub>M M2) f"  | 
| 47694 | 587  | 
by (subst distr_pair_swap)  | 
| 56996 | 588  | 
(auto simp: nn_integral_distr[OF measurable_pair_swap' f] intro!: nn_integral_cong)  | 
| 40859 | 589  | 
finally show ?thesis .  | 
590  | 
qed  | 
|
591  | 
||
592  | 
lemma (in pair_sigma_finite) Fubini:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
593  | 
assumes f: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
594  | 
shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M2) \<partial>M1)"  | 
| 56996 | 595  | 
unfolding nn_integral_snd[OF assms] M2.nn_integral_fst[OF assms] ..  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
596  | 
|
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
597  | 
lemma (in pair_sigma_finite) Fubini':  | 
| 
61424
 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 
haftmann 
parents: 
61169 
diff
changeset
 | 
598  | 
assumes f: "case_prod f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
599  | 
shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f x y \<partial>M1) \<partial>M2) = (\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f x y \<partial>M2) \<partial>M1)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
600  | 
using Fubini[OF f] by simp  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
601  | 
|
| 61808 | 602  | 
subsection \<open>Products on counting spaces, densities and distributions\<close>  | 
| 40859 | 603  | 
|
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
604  | 
lemma sigma_prod:  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
605  | 
assumes X_cover: "\<exists>E\<subseteq>A. countable E \<and> X = \<Union>E" and A: "A \<subseteq> Pow X"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
606  | 
assumes Y_cover: "\<exists>E\<subseteq>B. countable E \<and> Y = \<Union>E" and B: "B \<subseteq> Pow Y"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
607  | 
  shows "sigma X A \<Otimes>\<^sub>M sigma Y B = sigma (X \<times> Y) {a \<times> b | a b. a \<in> A \<and> b \<in> B}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
608  | 
(is "?P = ?S")  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
609  | 
proof (rule measure_eqI)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
610  | 
have [simp]: "snd \<in> X \<times> Y \<rightarrow> Y" "fst \<in> X \<times> Y \<rightarrow> X"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
611  | 
by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
612  | 
  let ?XY = "{{fst -` a \<inter> X \<times> Y | a. a \<in> A}, {snd -` b \<inter> X \<times> Y | b. b \<in> B}}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
613  | 
have "sets ?P =  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
614  | 
sets (\<Squnion>\<^sub>\<sigma> xy\<in>?XY. sigma (X \<times> Y) xy)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
615  | 
by (simp add: vimage_algebra_sigma sets_pair_eq_sets_fst_snd A B)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
616  | 
also have "\<dots> = sets (sigma (X \<times> Y) (\<Union>?XY))"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
617  | 
by (intro Sup_sigma_sigma arg_cong[where f=sets]) auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
618  | 
also have "\<dots> = sets ?S"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
619  | 
proof (intro arg_cong[where f=sets] sigma_eqI sigma_sets_eqI)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
620  | 
    show "\<Union>?XY \<subseteq> Pow (X \<times> Y)" "{a \<times> b |a b. a \<in> A \<and> b \<in> B} \<subseteq> Pow (X \<times> Y)"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
621  | 
using A B by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
622  | 
next  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
623  | 
    interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
624  | 
using A B by (intro sigma_algebra_sigma_sets) auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
625  | 
fix Z assume "Z \<in> \<Union>?XY"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
626  | 
    then show "Z \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
627  | 
proof safe  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
628  | 
fix a assume "a \<in> A"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
629  | 
from Y_cover obtain E where E: "E \<subseteq> B" "countable E" and "Y = \<Union>E"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
630  | 
by auto  | 
| 61808 | 631  | 
with \<open>a \<in> A\<close> A have eq: "fst -` a \<inter> X \<times> Y = (\<Union>e\<in>E. a \<times> e)"  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
632  | 
by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
633  | 
      show "fst -` a \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 61808 | 634  | 
using \<open>a \<in> A\<close> E unfolding eq by (auto intro!: XY.countable_UN')  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
635  | 
next  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
636  | 
fix b assume "b \<in> B"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
637  | 
from X_cover obtain E where E: "E \<subseteq> A" "countable E" and "X = \<Union>E"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
638  | 
by auto  | 
| 61808 | 639  | 
with \<open>b \<in> B\<close> B have eq: "snd -` b \<inter> X \<times> Y = (\<Union>e\<in>E. e \<times> b)"  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
640  | 
by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
641  | 
      show "snd -` b \<inter> X \<times> Y \<in> sigma_sets (X \<times> Y) {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 61808 | 642  | 
using \<open>b \<in> B\<close> E unfolding eq by (auto intro!: XY.countable_UN')  | 
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
643  | 
qed  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
644  | 
next  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
645  | 
    fix Z assume "Z \<in> {a \<times> b |a b. a \<in> A \<and> b \<in> B}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
646  | 
then obtain a b where "Z = a \<times> b" and ab: "a \<in> A" "b \<in> B"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
647  | 
by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
648  | 
then have Z: "Z = (fst -` a \<inter> X \<times> Y) \<inter> (snd -` b \<inter> X \<times> Y)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
649  | 
using A B by auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
650  | 
interpret XY: sigma_algebra "X \<times> Y" "sigma_sets (X \<times> Y) (\<Union>?XY)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
651  | 
by (intro sigma_algebra_sigma_sets) auto  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
652  | 
show "Z \<in> sigma_sets (X \<times> Y) (\<Union>?XY)"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
653  | 
unfolding Z by (rule XY.Int) (blast intro: ab)+  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
654  | 
qed  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
655  | 
finally show "sets ?P = sets ?S" .  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
656  | 
next  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
657  | 
interpret finite_measure "sigma X A" for X A  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
658  | 
proof qed (simp add: emeasure_sigma)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
659  | 
fix A assume "A \<in> sets ?P" then show "emeasure ?P A = emeasure ?S A"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
660  | 
by (simp add: emeasure_pair_measure_alt emeasure_sigma)  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
661  | 
qed  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
662  | 
|
| 
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
 | 
663  | 
lemma sigma_sets_pair_measure_generator_finite:  | 
| 38656 | 664  | 
assumes "finite A" and "finite B"  | 
| 47694 | 665  | 
  shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<subseteq> A \<and> b \<subseteq> B} = Pow (A \<times> B)"
 | 
| 40859 | 666  | 
(is "sigma_sets ?prod ?sets = _")  | 
| 38656 | 667  | 
proof safe  | 
668  | 
have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)  | 
|
669  | 
fix x assume subset: "x \<subseteq> A \<times> B"  | 
|
670  | 
hence "finite x" using fin by (rule finite_subset)  | 
|
| 40859 | 671  | 
from this subset show "x \<in> sigma_sets ?prod ?sets"  | 
| 38656 | 672  | 
proof (induct x)  | 
673  | 
case empty show ?case by (rule sigma_sets.Empty)  | 
|
674  | 
next  | 
|
675  | 
case (insert a x)  | 
|
| 47694 | 676  | 
    hence "{a} \<in> sigma_sets ?prod ?sets" by auto
 | 
| 38656 | 677  | 
moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto  | 
678  | 
ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)  | 
|
679  | 
qed  | 
|
680  | 
next  | 
|
681  | 
fix x a b  | 
|
| 40859 | 682  | 
assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x"  | 
| 38656 | 683  | 
from sigma_sets_into_sp[OF _ this(1)] this(2)  | 
| 40859 | 684  | 
show "a \<in> A" and "b \<in> B" by auto  | 
| 35833 | 685  | 
qed  | 
686  | 
||
| 
59088
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
687  | 
lemma borel_prod:  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
688  | 
  "(borel \<Otimes>\<^sub>M borel) = (borel :: ('a::second_countable_topology \<times> 'b::second_countable_topology) measure)"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
689  | 
(is "?P = ?B")  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
690  | 
proof -  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
691  | 
  have "?B = sigma UNIV {A \<times> B | A B. open A \<and> open B}"
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
692  | 
by (rule second_countable_borel_measurable[OF open_prod_generated])  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
693  | 
also have "\<dots> = ?P"  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
694  | 
unfolding borel_def  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
695  | 
    by (subst sigma_prod) (auto intro!: exI[of _ "{UNIV}"])
 | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
696  | 
finally show ?thesis ..  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
697  | 
qed  | 
| 
 
ff2bd4a14ddb
generalized (borel_)measurable_SUP/INF/lfp/gfp; tuned proofs for sigma-closure of product spaces
 
hoelzl 
parents: 
59048 
diff
changeset
 | 
698  | 
|
| 47694 | 699  | 
lemma pair_measure_count_space:  | 
700  | 
assumes A: "finite A" and B: "finite B"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
701  | 
shows "count_space A \<Otimes>\<^sub>M count_space B = count_space (A \<times> B)" (is "?P = ?C")  | 
| 47694 | 702  | 
proof (rule measure_eqI)  | 
703  | 
interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact  | 
|
704  | 
interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact  | 
|
| 61169 | 705  | 
interpret P: pair_sigma_finite "count_space A" "count_space B" ..  | 
| 47694 | 706  | 
show eq: "sets ?P = sets ?C"  | 
707  | 
by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B)  | 
|
708  | 
fix X assume X: "X \<in> sets ?P"  | 
|
709  | 
with eq have X_subset: "X \<subseteq> A \<times> B" by simp  | 
|
710  | 
with A B have fin_Pair: "\<And>x. finite (Pair x -` X)"  | 
|
711  | 
by (intro finite_subset[OF _ B]) auto  | 
|
712  | 
have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B)  | 
|
713  | 
show "emeasure ?P X = emeasure ?C X"  | 
|
| 49776 | 714  | 
apply (subst B.emeasure_pair_measure_alt[OF X])  | 
| 47694 | 715  | 
apply (subst emeasure_count_space)  | 
716  | 
using X_subset apply auto []  | 
|
717  | 
apply (simp add: fin_Pair emeasure_count_space X_subset fin_X)  | 
|
| 56996 | 718  | 
apply (subst nn_integral_count_space)  | 
| 47694 | 719  | 
using A apply simp  | 
| 
61609
 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 
paulson <lp15@cam.ac.uk> 
parents: 
61424 
diff
changeset
 | 
720  | 
apply (simp del: of_nat_setsum add: of_nat_setsum[symmetric])  | 
| 47694 | 721  | 
apply (subst card_gt_0_iff)  | 
722  | 
apply (simp add: fin_Pair)  | 
|
723  | 
apply (subst card_SigmaI[symmetric])  | 
|
724  | 
using A apply simp  | 
|
725  | 
using fin_Pair apply simp  | 
|
726  | 
using X_subset apply (auto intro!: arg_cong[where f=card])  | 
|
727  | 
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
 | 
728  | 
qed  | 
| 35833 | 729  | 
|
| 
59426
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
730  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
731  | 
lemma emeasure_prod_count_space:  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
732  | 
assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M M)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)")  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
733  | 
shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. indicator A (x, y) \<partial>?B \<partial>?A)"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
734  | 
by (rule emeasure_measure_of[OF pair_measure_def])  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
735  | 
(auto simp: countably_additive_def positive_def suminf_indicator nn_integral_nonneg A  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
736  | 
nn_integral_suminf[symmetric] dest: sets.sets_into_space)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
737  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
738  | 
lemma emeasure_prod_count_space_single[simp]: "emeasure (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) {x} = 1"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
739  | 
proof -  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
740  | 
  have [simp]: "\<And>a b x y. indicator {(a, b)} (x, y) = (indicator {a} x * indicator {b} y::ereal)"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
741  | 
by (auto split: split_indicator)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
742  | 
show ?thesis  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
743  | 
by (cases x)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
744  | 
(auto simp: emeasure_prod_count_space nn_integral_cmult sets_Pair nn_integral_max_0 one_ereal_def[symmetric])  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
745  | 
qed  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
746  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
747  | 
lemma emeasure_count_space_prod_eq:  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
748  | 
  fixes A :: "('a \<times> 'b) set"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
749  | 
assumes A: "A \<in> sets (count_space UNIV \<Otimes>\<^sub>M count_space UNIV)" (is "A \<in> sets (?A \<Otimes>\<^sub>M ?B)")  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
750  | 
shows "emeasure (?A \<Otimes>\<^sub>M ?B) A = emeasure (count_space UNIV) A"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
751  | 
proof -  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
752  | 
  { fix A :: "('a \<times> 'b) set" assume "countable A"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
753  | 
    then have "emeasure (?A \<Otimes>\<^sub>M ?B) (\<Union>a\<in>A. {a}) = (\<integral>\<^sup>+a. emeasure (?A \<Otimes>\<^sub>M ?B) {a} \<partial>count_space A)"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
754  | 
by (intro emeasure_UN_countable) (auto simp: sets_Pair disjoint_family_on_def)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
755  | 
also have "\<dots> = (\<integral>\<^sup>+a. indicator A a \<partial>count_space UNIV)"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
756  | 
by (subst nn_integral_count_space_indicator) auto  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
757  | 
finally have "emeasure (?A \<Otimes>\<^sub>M ?B) A = emeasure (count_space UNIV) A"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
758  | 
by simp }  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
759  | 
note * = this  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
760  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
761  | 
show ?thesis  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
762  | 
proof cases  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
763  | 
assume "finite A" then show ?thesis  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
764  | 
by (intro * countable_finite)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
765  | 
next  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
766  | 
assume "infinite A"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
767  | 
then obtain C where "countable C" and "infinite C" and "C \<subseteq> A"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
768  | 
by (auto dest: infinite_countable_subset')  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
769  | 
with A have "emeasure (?A \<Otimes>\<^sub>M ?B) C \<le> emeasure (?A \<Otimes>\<^sub>M ?B) A"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
770  | 
by (intro emeasure_mono) auto  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
771  | 
also have "emeasure (?A \<Otimes>\<^sub>M ?B) C = emeasure (count_space UNIV) C"  | 
| 61808 | 772  | 
using \<open>countable C\<close> by (rule *)  | 
| 
59426
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
773  | 
finally show ?thesis  | 
| 61808 | 774  | 
using \<open>infinite C\<close> \<open>infinite A\<close> by simp  | 
| 
59426
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
775  | 
qed  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
776  | 
qed  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
777  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
778  | 
lemma nn_intergal_count_space_prod_eq':  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
779  | 
assumes [simp]: "\<And>x. 0 \<le> f x"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
780  | 
shows "nn_integral (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
781  | 
(is "nn_integral ?P f = _")  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
782  | 
proof cases  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
783  | 
  assume cntbl: "countable {x. f x \<noteq> 0}"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
784  | 
  have [simp]: "\<And>x. ereal (real (card ({x} \<inter> {x. f x \<noteq> 0}))) = indicator {x. f x \<noteq> 0} x"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
785  | 
by (auto split: split_indicator)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
786  | 
  have [measurable]: "\<And>y. (\<lambda>x. indicator {y} x) \<in> borel_measurable ?P"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
787  | 
    by (rule measurable_discrete_difference[of "\<lambda>x. 0" _ borel "{y}" "\<lambda>x. indicator {y} x" for y])
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
788  | 
(auto intro: sets_Pair)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
789  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
790  | 
  have "(\<integral>\<^sup>+x. f x \<partial>?P) = (\<integral>\<^sup>+x. \<integral>\<^sup>+x'. f x * indicator {x} x' \<partial>count_space {x. f x \<noteq> 0} \<partial>?P)"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
791  | 
by (auto simp add: nn_integral_cmult nn_integral_indicator' intro!: nn_integral_cong split: split_indicator)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
792  | 
  also have "\<dots> = (\<integral>\<^sup>+x. \<integral>\<^sup>+x'. f x' * indicator {x'} x \<partial>count_space {x. f x \<noteq> 0} \<partial>?P)"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
793  | 
by (auto intro!: nn_integral_cong split: split_indicator)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
794  | 
  also have "\<dots> = (\<integral>\<^sup>+x'. \<integral>\<^sup>+x. f x' * indicator {x'} x \<partial>?P \<partial>count_space {x. f x \<noteq> 0})"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
795  | 
by (intro nn_integral_count_space_nn_integral cntbl) auto  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
796  | 
  also have "\<dots> = (\<integral>\<^sup>+x'. f x' \<partial>count_space {x. f x \<noteq> 0})"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
797  | 
by (intro nn_integral_cong) (auto simp: nn_integral_cmult sets_Pair)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
798  | 
finally show ?thesis  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
799  | 
by (auto simp add: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
800  | 
next  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
801  | 
  { fix x assume "f x \<noteq> 0"
 | 
| 61808 | 802  | 
with \<open>0 \<le> f x\<close> have "(\<exists>r. 0 < r \<and> f x = ereal r) \<or> f x = \<infinity>"  | 
| 
59426
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
803  | 
by (cases "f x") (auto simp: less_le)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
804  | 
then have "\<exists>n. ereal (1 / real (Suc n)) \<le> f x"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
805  | 
by (auto elim!: nat_approx_posE intro!: less_imp_le) }  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
806  | 
note * = this  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
807  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
808  | 
  assume cntbl: "uncountable {x. f x \<noteq> 0}"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
809  | 
  also have "{x. f x \<noteq> 0} = (\<Union>n. {x. 1/Suc n \<le> f x})"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
810  | 
using * by auto  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
811  | 
  finally obtain n where "infinite {x. 1/Suc n \<le> f x}"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
812  | 
by (meson countableI_type countable_UN uncountable_infinite)  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
813  | 
  then obtain C where C: "C \<subseteq> {x. 1/Suc n \<le> f x}" and "countable C" "infinite C"
 | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
814  | 
by (metis infinite_countable_subset')  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
815  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
816  | 
have [measurable]: "C \<in> sets ?P"  | 
| 61808 | 817  | 
using sets.countable[OF _ \<open>countable C\<close>, of ?P] by (auto simp: sets_Pair)  | 
| 
59426
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
818  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
819  | 
have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>?P) \<le> nn_integral ?P f"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
820  | 
using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric])  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
821  | 
moreover have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>?P) = \<infinity>"  | 
| 61808 | 822  | 
using \<open>infinite C\<close> by (simp add: nn_integral_cmult emeasure_count_space_prod_eq)  | 
| 
59426
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
823  | 
moreover have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>count_space UNIV) \<le> nn_integral (count_space UNIV) f"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
824  | 
using C by (intro nn_integral_mono) (auto split: split_indicator simp: zero_ereal_def[symmetric])  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
825  | 
moreover have "(\<integral>\<^sup>+x. ereal (1/Suc n) * indicator C x \<partial>count_space UNIV) = \<infinity>"  | 
| 61808 | 826  | 
using \<open>infinite C\<close> by (simp add: nn_integral_cmult)  | 
| 
59426
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
827  | 
ultimately show ?thesis  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
828  | 
by simp  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
829  | 
qed  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
830  | 
|
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
831  | 
lemma nn_intergal_count_space_prod_eq:  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
832  | 
"nn_integral (count_space UNIV \<Otimes>\<^sub>M count_space UNIV) f = nn_integral (count_space UNIV) f"  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
833  | 
by (subst (1 2) nn_integral_max_0[symmetric]) (auto intro!: nn_intergal_count_space_prod_eq')  | 
| 
 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 
hoelzl 
parents: 
59353 
diff
changeset
 | 
834  | 
|
| 47694 | 835  | 
lemma pair_measure_density:  | 
836  | 
assumes f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x"  | 
|
837  | 
assumes g: "g \<in> borel_measurable M2" "AE x in M2. 0 \<le> g x"  | 
|
| 50003 | 838  | 
assumes "sigma_finite_measure M2" "sigma_finite_measure (density M2 g)"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
839  | 
shows "density M1 f \<Otimes>\<^sub>M density M2 g = density (M1 \<Otimes>\<^sub>M M2) (\<lambda>(x,y). f x * g y)" (is "?L = ?R")  | 
| 47694 | 840  | 
proof (rule measure_eqI)  | 
841  | 
interpret M2: sigma_finite_measure M2 by fact  | 
|
842  | 
interpret D2: sigma_finite_measure "density M2 g" by fact  | 
|
843  | 
||
844  | 
fix A assume A: "A \<in> sets ?L"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
845  | 
with f g have "(\<integral>\<^sup>+ x. f x * \<integral>\<^sup>+ y. g y * indicator A (x, y) \<partial>M2 \<partial>M1) =  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
846  | 
(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f x * g y * indicator A (x, y) \<partial>M2 \<partial>M1)"  | 
| 56996 | 847  | 
by (intro nn_integral_cong_AE)  | 
848  | 
(auto simp add: nn_integral_cmult[symmetric] ac_simps)  | 
|
| 50003 | 849  | 
with A f g show "emeasure ?L A = emeasure ?R A"  | 
| 56996 | 850  | 
by (simp add: D2.emeasure_pair_measure emeasure_density nn_integral_density  | 
851  | 
M2.nn_integral_fst[symmetric]  | 
|
852  | 
cong: nn_integral_cong)  | 
|
| 47694 | 853  | 
qed simp  | 
854  | 
||
855  | 
lemma sigma_finite_measure_distr:  | 
|
856  | 
assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N"  | 
|
857  | 
shows "sigma_finite_measure M"  | 
|
| 40859 | 858  | 
proof -  | 
| 47694 | 859  | 
interpret sigma_finite_measure "distr M N f" by fact  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
860  | 
from sigma_finite_countable guess A .. note A = this  | 
| 47694 | 861  | 
show ?thesis  | 
| 
57447
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
862  | 
proof  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
863  | 
show "\<exists>A. countable A \<and> A \<subseteq> sets M \<and> \<Union>A = space M \<and> (\<forall>a\<in>A. emeasure M a \<noteq> \<infinity>)"  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
864  | 
using A f  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
865  | 
by (intro exI[of _ "(\<lambda>a. f -` a \<inter> space M) ` A"])  | 
| 
 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 
hoelzl 
parents: 
57235 
diff
changeset
 | 
866  | 
(auto simp: emeasure_distr set_eq_iff subset_eq intro: measurable_space)  | 
| 47694 | 867  | 
qed  | 
| 38656 | 868  | 
qed  | 
869  | 
||
| 47694 | 870  | 
lemma pair_measure_distr:  | 
871  | 
assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T"  | 
|
| 50003 | 872  | 
assumes "sigma_finite_measure (distr N T g)"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
873  | 
shows "distr M S f \<Otimes>\<^sub>M distr N T g = distr (M \<Otimes>\<^sub>M N) (S \<Otimes>\<^sub>M T) (\<lambda>(x, y). (f x, g y))" (is "?P = ?D")  | 
| 47694 | 874  | 
proof (rule measure_eqI)  | 
875  | 
interpret T: sigma_finite_measure "distr N T g" by fact  | 
|
876  | 
interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+  | 
|
| 50003 | 877  | 
|
| 47694 | 878  | 
fix A assume A: "A \<in> sets ?P"  | 
| 50003 | 879  | 
with f g show "emeasure ?P A = emeasure ?D A"  | 
880  | 
by (auto simp add: N.emeasure_pair_measure_alt space_pair_measure emeasure_distr  | 
|
| 56996 | 881  | 
T.emeasure_pair_measure_alt nn_integral_distr  | 
882  | 
intro!: nn_integral_cong arg_cong[where f="emeasure N"])  | 
|
| 50003 | 883  | 
qed simp  | 
| 39097 | 884  | 
|
| 50104 | 885  | 
lemma pair_measure_eqI:  | 
886  | 
assumes "sigma_finite_measure M1" "sigma_finite_measure M2"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
887  | 
assumes sets: "sets (M1 \<Otimes>\<^sub>M M2) = sets M"  | 
| 50104 | 888  | 
assumes emeasure: "\<And>A B. A \<in> sets M1 \<Longrightarrow> B \<in> sets M2 \<Longrightarrow> emeasure M1 A * emeasure M2 B = emeasure M (A \<times> B)"  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
889  | 
shows "M1 \<Otimes>\<^sub>M M2 = M"  | 
| 50104 | 890  | 
proof -  | 
891  | 
interpret M1: sigma_finite_measure M1 by fact  | 
|
892  | 
interpret M2: sigma_finite_measure M2 by fact  | 
|
| 61169 | 893  | 
interpret pair_sigma_finite M1 M2 ..  | 
| 50104 | 894  | 
  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
 | 
895  | 
  let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
 | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
896  | 
let ?P = "M1 \<Otimes>\<^sub>M M2"  | 
| 50104 | 897  | 
show ?thesis  | 
898  | 
proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])  | 
|
899  | 
show "?E \<subseteq> Pow (space ?P)"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
900  | 
using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure)  | 
| 50104 | 901  | 
show "sets ?P = sigma_sets (space ?P) ?E"  | 
902  | 
by (simp add: sets_pair_measure space_pair_measure)  | 
|
903  | 
then show "sets M = sigma_sets (space ?P) ?E"  | 
|
904  | 
using sets[symmetric] by simp  | 
|
905  | 
next  | 
|
906  | 
show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"  | 
|
907  | 
using F by (auto simp: space_pair_measure)  | 
|
908  | 
next  | 
|
909  | 
fix X assume "X \<in> ?E"  | 
|
910  | 
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  | 
|
911  | 
then have "emeasure ?P X = emeasure M1 A * emeasure M2 B"  | 
|
912  | 
by (simp add: M2.emeasure_pair_measure_Times)  | 
|
913  | 
also have "\<dots> = emeasure M (A \<times> B)"  | 
|
914  | 
using A B emeasure by auto  | 
|
915  | 
finally show "emeasure ?P X = emeasure M X"  | 
|
916  | 
by simp  | 
|
917  | 
qed  | 
|
918  | 
qed  | 
|
| 57025 | 919  | 
|
920  | 
lemma sets_pair_countable:  | 
|
921  | 
assumes "countable S1" "countable S2"  | 
|
922  | 
assumes M: "sets M = Pow S1" and N: "sets N = Pow S2"  | 
|
923  | 
shows "sets (M \<Otimes>\<^sub>M N) = Pow (S1 \<times> S2)"  | 
|
924  | 
proof auto  | 
|
925  | 
fix x a b assume x: "x \<in> sets (M \<Otimes>\<^sub>M N)" "(a, b) \<in> x"  | 
|
926  | 
from sets.sets_into_space[OF x(1)] x(2)  | 
|
927  | 
sets_eq_imp_space_eq[of N "count_space S2"] sets_eq_imp_space_eq[of M "count_space S1"] M N  | 
|
928  | 
show "a \<in> S1" "b \<in> S2"  | 
|
929  | 
by (auto simp: space_pair_measure)  | 
|
930  | 
next  | 
|
931  | 
fix X assume X: "X \<subseteq> S1 \<times> S2"  | 
|
932  | 
then have "countable X"  | 
|
| 61808 | 933  | 
by (metis countable_subset \<open>countable S1\<close> \<open>countable S2\<close> countable_SIGMA)  | 
| 57025 | 934  | 
  have "X = (\<Union>(a, b)\<in>X. {a} \<times> {b})" by auto
 | 
935  | 
also have "\<dots> \<in> sets (M \<Otimes>\<^sub>M N)"  | 
|
936  | 
using X  | 
|
| 61808 | 937  | 
by (safe intro!: sets.countable_UN' \<open>countable X\<close> subsetI pair_measureI) (auto simp: M N)  | 
| 57025 | 938  | 
finally show "X \<in> sets (M \<Otimes>\<^sub>M N)" .  | 
939  | 
qed  | 
|
940  | 
||
941  | 
lemma pair_measure_countable:  | 
|
942  | 
assumes "countable S1" "countable S2"  | 
|
943  | 
shows "count_space S1 \<Otimes>\<^sub>M count_space S2 = count_space (S1 \<times> S2)"  | 
|
944  | 
proof (rule pair_measure_eqI)  | 
|
945  | 
show "sigma_finite_measure (count_space S1)" "sigma_finite_measure (count_space S2)"  | 
|
946  | 
using assms by (auto intro!: sigma_finite_measure_count_space_countable)  | 
|
947  | 
show "sets (count_space S1 \<Otimes>\<^sub>M count_space S2) = sets (count_space (S1 \<times> S2))"  | 
|
948  | 
by (subst sets_pair_countable[OF assms]) auto  | 
|
949  | 
next  | 
|
950  | 
fix A B assume "A \<in> sets (count_space S1)" "B \<in> sets (count_space S2)"  | 
|
951  | 
then show "emeasure (count_space S1) A * emeasure (count_space S2) B =  | 
|
952  | 
emeasure (count_space (S1 \<times> S2)) (A \<times> B)"  | 
|
953  | 
by (subst (1 2 3) emeasure_count_space) (auto simp: finite_cartesian_product_iff)  | 
|
954  | 
qed  | 
|
| 50104 | 955  | 
|
| 
59489
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
956  | 
lemma nn_integral_fst_count_space':  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
957  | 
assumes nonneg: "\<And>xy. 0 \<le> f xy"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
958  | 
shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
959  | 
(is "?lhs = ?rhs")  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
960  | 
proof(cases)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
961  | 
  assume *: "countable {xy. f xy \<noteq> 0}"
 | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
962  | 
  let ?A = "fst ` {xy. f xy \<noteq> 0}"
 | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
963  | 
  let ?B = "snd ` {xy. f xy \<noteq> 0}"
 | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
964  | 
from * have [simp]: "countable ?A" "countable ?B" by(rule countable_image)+  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
965  | 
from nonneg have f_neq_0: "\<And>xy. f xy \<noteq> 0 \<longleftrightarrow> f xy > 0"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
966  | 
by(auto simp add: order.order_iff_strict)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
967  | 
have "?lhs = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space ?A)"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
968  | 
by(rule nn_integral_count_space_eq)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
969  | 
(auto simp add: f_neq_0 nn_integral_0_iff_AE AE_count_space not_le intro: rev_image_eqI)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
970  | 
also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space ?B \<partial>count_space ?A)"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
971  | 
by(intro nn_integral_count_space_eq nn_integral_cong)(auto intro: rev_image_eqI)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
972  | 
also have "\<dots> = (\<integral>\<^sup>+ xy. f xy \<partial>count_space (?A \<times> ?B))"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
973  | 
by(subst sigma_finite_measure.nn_integral_fst)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
974  | 
(simp_all add: sigma_finite_measure_count_space_countable pair_measure_countable)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
975  | 
also have "\<dots> = ?rhs"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
976  | 
by(rule nn_integral_count_space_eq)(auto intro: rev_image_eqI)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
977  | 
finally show ?thesis .  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
978  | 
next  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
979  | 
  { fix xy assume "f xy \<noteq> 0"
 | 
| 61808 | 980  | 
with \<open>0 \<le> f xy\<close> have "(\<exists>r. 0 < r \<and> f xy = ereal r) \<or> f xy = \<infinity>"  | 
| 
59489
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
981  | 
by (cases "f xy") (auto simp: less_le)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
982  | 
then have "\<exists>n. ereal (1 / real (Suc n)) \<le> f xy"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
983  | 
by (auto elim!: nat_approx_posE intro!: less_imp_le) }  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
984  | 
note * = this  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
985  | 
|
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
986  | 
  assume cntbl: "uncountable {xy. f xy \<noteq> 0}"
 | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
987  | 
  also have "{xy. f xy \<noteq> 0} = (\<Union>n. {xy. 1/Suc n \<le> f xy})"
 | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
988  | 
using * by auto  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
989  | 
  finally obtain n where "infinite {xy. 1/Suc n \<le> f xy}"
 | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
990  | 
by (meson countableI_type countable_UN uncountable_infinite)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
991  | 
  then obtain C where C: "C \<subseteq> {xy. 1/Suc n \<le> f xy}" and "countable C" "infinite C"
 | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
992  | 
by (metis infinite_countable_subset')  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
993  | 
|
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
994  | 
have "\<infinity> = (\<integral>\<^sup>+ xy. ereal (1 / Suc n) * indicator C xy \<partial>count_space UNIV)"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
995  | 
using \<open>infinite C\<close> by(simp add: nn_integral_cmult)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
996  | 
also have "\<dots> \<le> ?rhs" using C  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
997  | 
by(intro nn_integral_mono)(auto split: split_indicator simp add: nonneg)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
998  | 
finally have "?rhs = \<infinity>" by simp  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
999  | 
moreover have "?lhs = \<infinity>"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1000  | 
proof(cases "finite (fst ` C)")  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1001  | 
case True  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1002  | 
then obtain x C' where x: "x \<in> fst ` C"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1003  | 
      and C': "C' = fst -` {x} \<inter> C"
 | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1004  | 
and "infinite C'"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1005  | 
using \<open>infinite C\<close> by(auto elim!: inf_img_fin_domE')  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1006  | 
    from x C C' have **: "C' \<subseteq> {xy. 1 / Suc n \<le> f xy}" by auto
 | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1007  | 
|
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1008  | 
from C' \<open>infinite C'\<close> have "infinite (snd ` C')"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1009  | 
by(auto dest!: finite_imageD simp add: inj_on_def)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1010  | 
then have "\<infinity> = (\<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator (snd ` C') y \<partial>count_space UNIV)"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1011  | 
by(simp add: nn_integral_cmult)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1012  | 
also have "\<dots> = (\<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV)"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1013  | 
by(rule nn_integral_cong)(force split: split_indicator intro: rev_image_eqI simp add: C')  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1014  | 
    also have "\<dots> = (\<integral>\<^sup>+ x'. (\<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV) * indicator {x} x' \<partial>count_space UNIV)"
 | 
| 62083 | 1015  | 
by(simp add: one_ereal_def[symmetric] nn_integral_nonneg max_def)  | 
| 
59489
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1016  | 
also have "\<dots> \<le> (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C' (x, y) \<partial>count_space UNIV \<partial>count_space UNIV)"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1017  | 
by(rule nn_integral_mono)(simp split: split_indicator add: nn_integral_nonneg)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1018  | 
also have "\<dots> \<le> ?lhs" using **  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1019  | 
by(intro nn_integral_mono)(auto split: split_indicator simp add: nonneg)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1020  | 
finally show ?thesis by simp  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1021  | 
next  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1022  | 
case False  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1023  | 
def C' \<equiv> "fst ` C"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1024  | 
have "\<infinity> = \<integral>\<^sup>+ x. ereal (1 / Suc n) * indicator C' x \<partial>count_space UNIV"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1025  | 
using C'_def False by(simp add: nn_integral_cmult)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1026  | 
    also have "\<dots> = \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C' x * indicator {SOME y. (x, y) \<in> C} y \<partial>count_space UNIV \<partial>count_space UNIV"
 | 
| 62083 | 1027  | 
by(auto simp add: one_ereal_def[symmetric] max_def intro: nn_integral_cong)  | 
| 
59489
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1028  | 
also have "\<dots> \<le> \<integral>\<^sup>+ x. \<integral>\<^sup>+ y. ereal (1 / Suc n) * indicator C (x, y) \<partial>count_space UNIV \<partial>count_space UNIV"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1029  | 
by(intro nn_integral_mono)(auto simp add: C'_def split: split_indicator intro: someI)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1030  | 
also have "\<dots> \<le> ?lhs" using C  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1031  | 
by(intro nn_integral_mono)(auto split: split_indicator simp add: nonneg)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1032  | 
finally show ?thesis by simp  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1033  | 
qed  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1034  | 
ultimately show ?thesis by simp  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1035  | 
qed  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1036  | 
|
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1037  | 
lemma nn_integral_fst_count_space:  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1038  | 
"(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1039  | 
by(subst (2 3) nn_integral_max_0[symmetric])(rule nn_integral_fst_count_space', simp)  | 
| 
 
fd5d23cc0e97
nn_integral can be split over arbitrary product count_spaces
 
Andreas Lochbihler 
parents: 
59426 
diff
changeset
 | 
1040  | 
|
| 
59491
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1041  | 
lemma nn_integral_snd_count_space:  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1042  | 
"(\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. f (x, y) \<partial>count_space UNIV \<partial>count_space UNIV) = integral\<^sup>N (count_space UNIV) f"  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1043  | 
(is "?lhs = ?rhs")  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1044  | 
proof -  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1045  | 
have "?lhs = (\<integral>\<^sup>+ y. \<integral>\<^sup>+ x. (\<lambda>(y, x). f (x, y)) (y, x) \<partial>count_space UNIV \<partial>count_space UNIV)"  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1046  | 
by(simp)  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1047  | 
also have "\<dots> = \<integral>\<^sup>+ yx. (\<lambda>(y, x). f (x, y)) yx \<partial>count_space UNIV"  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1048  | 
by(rule nn_integral_fst_count_space)  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1049  | 
also have "\<dots> = \<integral>\<^sup>+ xy. f xy \<partial>count_space ((\<lambda>(x, y). (y, x)) ` UNIV)"  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1050  | 
by(subst nn_integral_bij_count_space[OF inj_on_imp_bij_betw, symmetric])  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1051  | 
(simp_all add: inj_on_def split_def)  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1052  | 
also have "\<dots> = ?rhs" by(rule nn_integral_count_space_eq) auto  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1053  | 
finally show ?thesis .  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1054  | 
qed  | 
| 
 
40f570f9a284
add another lemma to split nn_integral over product count_space
 
Andreas Lochbihler 
parents: 
59489 
diff
changeset
 | 
1055  | 
|
| 60066 | 1056  | 
lemma measurable_pair_measure_countable1:  | 
1057  | 
assumes "countable A"  | 
|
1058  | 
and [measurable]: "\<And>x. x \<in> A \<Longrightarrow> (\<lambda>y. f (x, y)) \<in> measurable N K"  | 
|
1059  | 
shows "f \<in> measurable (count_space A \<Otimes>\<^sub>M N) K"  | 
|
1060  | 
using _ _ assms(1)  | 
|
1061  | 
by(rule measurable_compose_countable'[where f="\<lambda>a b. f (a, snd b)" and g=fst and I=A, simplified])simp_all  | 
|
1062  | 
||
| 61808 | 1063  | 
subsection \<open>Product of Borel spaces\<close>  | 
| 
57235
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1064  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1065  | 
lemma borel_Times:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1066  | 
fixes A :: "'a::topological_space set" and B :: "'b::topological_space set"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1067  | 
assumes A: "A \<in> sets borel" and B: "B \<in> sets borel"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1068  | 
shows "A \<times> B \<in> sets borel"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1069  | 
proof -  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1070  | 
have "A \<times> B = (A\<times>UNIV) \<inter> (UNIV \<times> B)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1071  | 
by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1072  | 
moreover  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1073  | 
  { have "A \<in> sigma_sets UNIV {S. open S}" using A by (simp add: sets_borel)
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1074  | 
then have "A\<times>UNIV \<in> sets borel"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1075  | 
proof (induct A)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1076  | 
case (Basic S) then show ?case  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1077  | 
by (auto intro!: borel_open open_Times)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1078  | 
next  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1079  | 
case (Compl A)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1080  | 
moreover have *: "(UNIV - A) \<times> UNIV = UNIV - (A \<times> UNIV)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1081  | 
by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1082  | 
ultimately show ?case  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1083  | 
unfolding * by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1084  | 
next  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1085  | 
case (Union A)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1086  | 
moreover have *: "(UNION UNIV A) \<times> UNIV = UNION UNIV (\<lambda>i. A i \<times> UNIV)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1087  | 
by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1088  | 
ultimately show ?case  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1089  | 
unfolding * by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1090  | 
qed simp }  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1091  | 
moreover  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1092  | 
  { have "B \<in> sigma_sets UNIV {S. open S}" using B by (simp add: sets_borel)
 | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1093  | 
then have "UNIV\<times>B \<in> sets borel"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1094  | 
proof (induct B)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1095  | 
case (Basic S) then show ?case  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1096  | 
by (auto intro!: borel_open open_Times)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1097  | 
next  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1098  | 
case (Compl B)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1099  | 
moreover have *: "UNIV \<times> (UNIV - B) = UNIV - (UNIV \<times> B)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1100  | 
by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1101  | 
ultimately show ?case  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1102  | 
unfolding * by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1103  | 
next  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1104  | 
case (Union B)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1105  | 
moreover have *: "UNIV \<times> (UNION UNIV B) = UNION UNIV (\<lambda>i. UNIV \<times> B i)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1106  | 
by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1107  | 
ultimately show ?case  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1108  | 
unfolding * by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1109  | 
qed simp }  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1110  | 
ultimately show ?thesis  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1111  | 
by auto  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1112  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1113  | 
|
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1114  | 
lemma finite_measure_pair_measure:  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1115  | 
assumes "finite_measure M" "finite_measure N"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1116  | 
shows "finite_measure (N \<Otimes>\<^sub>M M)"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1117  | 
proof (rule finite_measureI)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1118  | 
interpret M: finite_measure M by fact  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1119  | 
interpret N: finite_measure N by fact  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1120  | 
show "emeasure (N \<Otimes>\<^sub>M M) (space (N \<Otimes>\<^sub>M M)) \<noteq> \<infinity>"  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1121  | 
by (auto simp: space_pair_measure M.emeasure_pair_measure_Times)  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1122  | 
qed  | 
| 
 
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
 
hoelzl 
parents: 
57025 
diff
changeset
 | 
1123  | 
|
| 62083 | 1124  | 
end  |