| author | blanchet | 
| Sun, 16 Feb 2014 18:39:41 +0100 | |
| changeset 55519 | 8a54bf4a92ca | 
| parent 54863 | 82acc20ded73 | 
| child 56993 | e5366291d6aa | 
| permissions | -rw-r--r-- | 
| 
42146
 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 
hoelzl 
parents: 
42067 
diff
changeset
 | 
1  | 
(* Title: HOL/Probability/Binary_Product_Measure.thy  | 
| 42067 | 2  | 
Author: Johannes Hölzl, TU München  | 
3  | 
*)  | 
|
4  | 
||
| 
42146
 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 
hoelzl 
parents: 
42067 
diff
changeset
 | 
5  | 
header {*Binary product measures*}
 | 
| 42067 | 6  | 
|
| 
42146
 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 
hoelzl 
parents: 
42067 
diff
changeset
 | 
7  | 
theory Binary_Product_Measure  | 
| 38656 | 8  | 
imports Lebesgue_Integration  | 
| 35833 | 9  | 
begin  | 
10  | 
||
| 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  | 
||
17  | 
section "Binary products"  | 
|
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  | 
|
32  | 
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
 | 
33  | 
  "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
 | 
34  | 
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
 | 
35  | 
by (rule sets_measure_of)  | 
| 41095 | 36  | 
|
| 49776 | 37  | 
lemma sets_pair_measure_cong[cong]:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
38  | 
"sets M1 = sets M1' \<Longrightarrow> sets M2 = sets M2' \<Longrightarrow> sets (M1 \<Otimes>\<^sub>M M2) = sets (M1' \<Otimes>\<^sub>M M2')"  | 
| 49776 | 39  | 
unfolding sets_pair_measure by (simp cong: sets_eq_imp_space_eq)  | 
40  | 
||
| 50003 | 41  | 
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
 | 
42  | 
"x \<in> sets A \<Longrightarrow> y \<in> sets B \<Longrightarrow> x \<times> y \<in> sets (A \<Otimes>\<^sub>M B)"  | 
| 47694 | 43  | 
by (auto simp: sets_pair_measure)  | 
| 41095 | 44  | 
|
| 47694 | 45  | 
lemma measurable_pair_measureI:  | 
46  | 
assumes 1: "f \<in> space M \<rightarrow> space M1 \<times> space M2"  | 
|
47  | 
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
 | 
48  | 
shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"  | 
| 47694 | 49  | 
unfolding pair_measure_def using 1 2  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
50  | 
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
 | 
51  | 
|
| 50003 | 52  | 
lemma measurable_split_replace[measurable (raw)]:  | 
53  | 
"(\<lambda>x. f x (fst (g x)) (snd (g x))) \<in> measurable M N \<Longrightarrow> (\<lambda>x. split (f x) (g x)) \<in> measurable M N"  | 
|
54  | 
unfolding split_beta' .  | 
|
55  | 
||
56  | 
lemma measurable_Pair[measurable (raw)]:  | 
|
| 49776 | 57  | 
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
 | 
58  | 
shows "(\<lambda>x. (f x, g x)) \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"  | 
| 49776 | 59  | 
proof (rule measurable_pair_measureI)  | 
60  | 
show "(\<lambda>x. (f x, g x)) \<in> space M \<rightarrow> space M1 \<times> space M2"  | 
|
61  | 
using f g by (auto simp: measurable_def)  | 
|
62  | 
fix A B assume *: "A \<in> sets M1" "B \<in> sets M2"  | 
|
63  | 
have "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)"  | 
|
64  | 
by auto  | 
|
65  | 
also have "\<dots> \<in> sets M"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
66  | 
by (rule sets.Int) (auto intro!: measurable_sets * f g)  | 
| 49776 | 67  | 
finally show "(\<lambda>x. (f x, g x)) -` (A \<times> B) \<inter> space M \<in> sets M" .  | 
68  | 
qed  | 
|
69  | 
||
| 50003 | 70  | 
lemma measurable_Pair_compose_split[measurable_dest]:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
71  | 
assumes f: "split f \<in> measurable (M1 \<Otimes>\<^sub>M M2) N"  | 
| 50003 | 72  | 
assumes g: "g \<in> measurable M M1" and h: "h \<in> measurable M M2"  | 
73  | 
shows "(\<lambda>x. f (g x) (h x)) \<in> measurable M N"  | 
|
74  | 
using measurable_compose[OF measurable_Pair f, OF g h] by simp  | 
|
75  | 
||
| 49776 | 76  | 
lemma measurable_pair:  | 
77  | 
assumes "(fst \<circ> f) \<in> measurable M M1" "(snd \<circ> f) \<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
 | 
78  | 
shows "f \<in> measurable M (M1 \<Otimes>\<^sub>M M2)"  | 
| 49776 | 79  | 
using measurable_Pair[OF assms] by simp  | 
80  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
81  | 
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
 | 
82  | 
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
 | 
83  | 
measurable_def)  | 
| 40859 | 84  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
85  | 
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
 | 
86  | 
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
 | 
87  | 
measurable_def)  | 
| 47694 | 88  | 
|
| 50003 | 89  | 
lemma  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
90  | 
assumes f[measurable]: "f \<in> measurable M (N \<Otimes>\<^sub>M P)"  | 
| 50003 | 91  | 
shows measurable_fst': "(\<lambda>x. fst (f x)) \<in> measurable M N"  | 
92  | 
and measurable_snd': "(\<lambda>x. snd (f x)) \<in> measurable M P"  | 
|
93  | 
by simp_all  | 
|
| 40859 | 94  | 
|
| 50003 | 95  | 
lemma  | 
96  | 
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
 | 
97  | 
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
 | 
98  | 
and measurable_snd'': "(\<lambda>x. f (snd x)) \<in> measurable (P \<Otimes>\<^sub>M M) N"  | 
| 50003 | 99  | 
by simp_all  | 
| 47694 | 100  | 
|
101  | 
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
 | 
102  | 
"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 | 103  | 
by (auto intro: measurable_pair[of f M M1 M2])  | 
| 40859 | 104  | 
|
| 49776 | 105  | 
lemma measurable_split_conv:  | 
106  | 
"(\<lambda>(x, y). f x y) \<in> measurable A B \<longleftrightarrow> (\<lambda>x. f (fst x) (snd x)) \<in> measurable A B"  | 
|
107  | 
by (intro arg_cong2[where f="op \<in>"]) auto  | 
|
| 40859 | 108  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
109  | 
lemma measurable_pair_swap': "(\<lambda>(x,y). (y, x)) \<in> measurable (M1 \<Otimes>\<^sub>M M2) (M2 \<Otimes>\<^sub>M M1)"  | 
| 49776 | 110  | 
by (auto intro!: measurable_Pair simp: measurable_split_conv)  | 
| 47694 | 111  | 
|
112  | 
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
 | 
113  | 
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 | 114  | 
using measurable_comp[OF measurable_Pair f] by (auto simp: measurable_split_conv comp_def)  | 
| 40859 | 115  | 
|
| 47694 | 116  | 
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
 | 
117  | 
"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 | 118  | 
by (auto dest: measurable_pair_swap)  | 
| 49776 | 119  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
120  | 
lemma measurable_Pair1': "x \<in> space M1 \<Longrightarrow> Pair x \<in> measurable M2 (M1 \<Otimes>\<^sub>M M2)"  | 
| 50003 | 121  | 
by simp  | 
| 40859 | 122  | 
|
| 50003 | 123  | 
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
 | 
124  | 
assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "Pair x -` A \<in> sets M2"  | 
| 40859 | 125  | 
proof -  | 
| 47694 | 126  | 
  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
 | 
127  | 
using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)  | 
| 47694 | 128  | 
also have "\<dots> \<in> sets M2"  | 
129  | 
using A by (auto simp add: measurable_Pair1' intro!: measurable_sets split: split_if_asm)  | 
|
130  | 
finally show ?thesis .  | 
|
| 40859 | 131  | 
qed  | 
132  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
133  | 
lemma measurable_Pair2': "y \<in> space M2 \<Longrightarrow> (\<lambda>x. (x, y)) \<in> measurable M1 (M1 \<Otimes>\<^sub>M M2)"  | 
| 49776 | 134  | 
by (auto intro!: measurable_Pair)  | 
| 40859 | 135  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
136  | 
lemma sets_Pair2: assumes A: "A \<in> sets (M1 \<Otimes>\<^sub>M M2)" shows "(\<lambda>x. (x, y)) -` A \<in> sets M1"  | 
| 47694 | 137  | 
proof -  | 
138  | 
  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
 | 
139  | 
using A[THEN sets.sets_into_space] by (auto simp: space_pair_measure)  | 
| 47694 | 140  | 
also have "\<dots> \<in> sets M1"  | 
141  | 
using A by (auto simp add: measurable_Pair2' intro!: measurable_sets split: split_if_asm)  | 
|
142  | 
finally show ?thesis .  | 
|
| 40859 | 143  | 
qed  | 
144  | 
||
| 47694 | 145  | 
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
 | 
146  | 
assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and x: "x \<in> space M1"  | 
| 47694 | 147  | 
shows "(\<lambda>y. f (x, y)) \<in> measurable M2 M"  | 
148  | 
using measurable_comp[OF measurable_Pair1' f, OF x]  | 
|
149  | 
by (simp add: comp_def)  | 
|
150  | 
||
151  | 
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
 | 
152  | 
assumes f: "f \<in> measurable (M1 \<Otimes>\<^sub>M M2) M" and y: "y \<in> space M2"  | 
| 40859 | 153  | 
shows "(\<lambda>x. f (x, y)) \<in> measurable M1 M"  | 
| 47694 | 154  | 
using measurable_comp[OF measurable_Pair2' f, OF y]  | 
155  | 
by (simp add: comp_def)  | 
|
| 40859 | 156  | 
|
| 47694 | 157  | 
lemma Int_stable_pair_measure_generator: "Int_stable {a \<times> b | a b. a \<in> sets A \<and> b \<in> sets B}"
 | 
| 40859 | 158  | 
unfolding Int_stable_def  | 
| 47694 | 159  | 
by safe (auto simp add: times_Int_times)  | 
| 40859 | 160  | 
|
| 50003 | 161  | 
lemma disjoint_family_vimageI: "disjoint_family F \<Longrightarrow> disjoint_family (\<lambda>i. f -` F i)"  | 
162  | 
by (auto simp: disjoint_family_on_def)  | 
|
163  | 
||
| 49776 | 164  | 
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
 | 
165  | 
assumes [measurable]: "Q \<in> sets (N \<Otimes>\<^sub>M M)"  | 
| 49776 | 166  | 
shows "(\<lambda>x. emeasure M (Pair x -` Q)) \<in> borel_measurable N"  | 
| 40859 | 167  | 
(is "?s Q \<in> _")  | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
168  | 
using Int_stable_pair_measure_generator pair_measure_closed assms  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
169  | 
unfolding sets_pair_measure  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
170  | 
proof (induct rule: sigma_sets_induct_disjoint)  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
171  | 
case (compl A)  | 
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
172  | 
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
 | 
173  | 
(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
 | 
174  | 
unfolding sets_pair_measure[symmetric]  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
175  | 
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
 | 
176  | 
with compl sets.top show ?case  | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
177  | 
by (auto intro!: measurable_If simp: space_pair_measure)  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
178  | 
next  | 
| 
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
179  | 
case (union F)  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
180  | 
then have "\<And>x. emeasure M (Pair x -` (\<Union>i. F i)) = (\<Sum>i. ?s (F i) x)"  | 
| 50003 | 181  | 
by (simp add: suminf_emeasure disjoint_family_vimageI subset_eq vimage_UN sets_pair_measure[symmetric])  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
182  | 
with union show ?case  | 
| 50003 | 183  | 
unfolding sets_pair_measure[symmetric] by simp  | 
| 
49789
 
e0a4cb91a8a9
add induction rule for intersection-stable sigma-sets
 
hoelzl 
parents: 
49784 
diff
changeset
 | 
184  | 
qed (auto simp add: if_distrib Int_def[symmetric] intro!: measurable_If)  | 
| 49776 | 185  | 
|
186  | 
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
 | 
187  | 
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 | 188  | 
proof -  | 
189  | 
from sigma_finite_disjoint guess F . note F = this  | 
|
190  | 
then have F_sets: "\<And>i. F i \<in> sets M" by auto  | 
|
191  | 
let ?C = "\<lambda>x i. F i \<inter> Pair x -` Q"  | 
|
192  | 
  { fix i
 | 
|
193  | 
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
 | 
194  | 
using F sets.sets_into_space by auto  | 
| 49776 | 195  | 
let ?R = "density M (indicator (F i))"  | 
196  | 
have "finite_measure ?R"  | 
|
197  | 
using F by (intro finite_measureI) (auto simp: emeasure_restricted subset_eq)  | 
|
198  | 
then have "(\<lambda>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))) \<in> borel_measurable N"  | 
|
199  | 
by (rule finite_measure.finite_measure_cut_measurable) (auto intro: Q)  | 
|
200  | 
moreover have "\<And>x. emeasure ?R (Pair x -` (space N \<times> space ?R \<inter> Q))  | 
|
201  | 
= emeasure M (F i \<inter> Pair x -` (space N \<times> space ?R \<inter> Q))"  | 
|
202  | 
using Q F_sets by (intro emeasure_restricted) (auto intro: sets_Pair1)  | 
|
203  | 
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
 | 
204  | 
using sets.sets_into_space[OF Q] by (auto simp: space_pair_measure)  | 
| 49776 | 205  | 
ultimately have "(\<lambda>x. emeasure M (?C x i)) \<in> borel_measurable N"  | 
206  | 
by simp }  | 
|
207  | 
moreover  | 
|
208  | 
  { fix x
 | 
|
209  | 
have "(\<Sum>i. emeasure M (?C x i)) = emeasure M (\<Union>i. ?C x i)"  | 
|
210  | 
proof (intro suminf_emeasure)  | 
|
211  | 
show "range (?C x) \<subseteq> sets M"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
212  | 
using F `Q \<in> sets (N \<Otimes>\<^sub>M M)` by (auto intro!: sets_Pair1)  | 
| 49776 | 213  | 
have "disjoint_family F" using F by auto  | 
214  | 
show "disjoint_family (?C x)"  | 
|
215  | 
by (rule disjoint_family_on_bisimulation[OF `disjoint_family F`]) auto  | 
|
216  | 
qed  | 
|
217  | 
also have "(\<Union>i. ?C x i) = Pair x -` Q"  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
218  | 
using F sets.sets_into_space[OF `Q \<in> sets (N \<Otimes>\<^sub>M M)`]  | 
| 49776 | 219  | 
by (auto simp: space_pair_measure)  | 
220  | 
finally have "emeasure M (Pair x -` Q) = (\<Sum>i. emeasure M (?C x i))"  | 
|
221  | 
by simp }  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
222  | 
ultimately show ?thesis using `Q \<in> sets (N \<Otimes>\<^sub>M M)` F_sets  | 
| 49776 | 223  | 
by auto  | 
224  | 
qed  | 
|
225  | 
||
| 50003 | 226  | 
lemma (in sigma_finite_measure) measurable_emeasure[measurable (raw)]:  | 
227  | 
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
 | 
228  | 
  assumes A: "{x\<in>space (N \<Otimes>\<^sub>M M). snd x \<in> A (fst x)} \<in> sets (N \<Otimes>\<^sub>M M)"
 | 
| 50003 | 229  | 
shows "(\<lambda>x. emeasure M (A x)) \<in> borel_measurable N"  | 
230  | 
proof -  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
231  | 
  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 | 232  | 
by (auto simp: space_pair_measure)  | 
233  | 
with measurable_emeasure_Pair[OF A] show ?thesis  | 
|
234  | 
by (auto cong: measurable_cong)  | 
|
235  | 
qed  | 
|
236  | 
||
| 49776 | 237  | 
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
 | 
238  | 
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
 | 
239  | 
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 | 240  | 
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
 | 
241  | 
show "positive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"  | 
| 49776 | 242  | 
by (auto simp: positive_def positive_integral_positive)  | 
243  | 
have eq[simp]: "\<And>A x y. indicator A (x, y) = indicator (Pair x -` A) y"  | 
|
244  | 
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
 | 
245  | 
show "countably_additive (sets (N \<Otimes>\<^sub>M M)) ?\<mu>"  | 
| 49776 | 246  | 
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
 | 
247  | 
    fix F :: "nat \<Rightarrow> ('b \<times> 'a) set" assume F: "range F \<subseteq> sets (N \<Otimes>\<^sub>M M)" "disjoint_family F"
 | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
248  | 
from F have *: "\<And>i. F i \<in> sets (N \<Otimes>\<^sub>M M)" "(\<Union>i. F i) \<in> sets (N \<Otimes>\<^sub>M M)" by auto  | 
| 49776 | 249  | 
moreover from F have "\<And>i. (\<lambda>x. emeasure M (Pair x -` F i)) \<in> borel_measurable N"  | 
250  | 
by (intro measurable_emeasure_Pair) auto  | 
|
251  | 
moreover have "\<And>x. disjoint_family (\<lambda>i. Pair x -` F i)"  | 
|
252  | 
by (intro disjoint_family_on_bisimulation[OF F(2)]) auto  | 
|
253  | 
moreover have "\<And>x. range (\<lambda>i. Pair x -` F i) \<subseteq> sets M"  | 
|
254  | 
using F by (auto simp: sets_Pair1)  | 
|
255  | 
ultimately show "(\<Sum>n. ?\<mu> (F n)) = ?\<mu> (\<Union>i. F i)"  | 
|
256  | 
by (auto simp add: vimage_UN positive_integral_suminf[symmetric] suminf_emeasure subset_eq emeasure_nonneg sets_Pair1  | 
|
257  | 
intro!: positive_integral_cong positive_integral_indicator[symmetric])  | 
|
258  | 
qed  | 
|
259  | 
  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
 | 
260  | 
using sets.space_closed[of N] sets.space_closed[of M] by auto  | 
| 49776 | 261  | 
qed fact  | 
262  | 
||
263  | 
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
 | 
264  | 
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
 | 
265  | 
shows "emeasure (N \<Otimes>\<^sub>M M) X = (\<integral>\<^sup>+x. emeasure M (Pair x -` X) \<partial>N)"  | 
| 49776 | 266  | 
proof -  | 
267  | 
have [simp]: "\<And>x y. indicator X (x, y) = indicator (Pair x -` X) y"  | 
|
268  | 
by (auto simp: indicator_def)  | 
|
269  | 
show ?thesis  | 
|
270  | 
using X by (auto intro!: positive_integral_cong simp: emeasure_pair_measure sets_Pair1)  | 
|
271  | 
qed  | 
|
272  | 
||
273  | 
lemma (in sigma_finite_measure) emeasure_pair_measure_Times:  | 
|
274  | 
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
 | 
275  | 
shows "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = emeasure N A * emeasure M B"  | 
| 49776 | 276  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
277  | 
have "emeasure (N \<Otimes>\<^sub>M M) (A \<times> B) = (\<integral>\<^sup>+x. emeasure M B * indicator A x \<partial>N)"  | 
| 49776 | 278  | 
using A B by (auto intro!: positive_integral_cong simp: emeasure_pair_measure_alt)  | 
279  | 
also have "\<dots> = emeasure M B * emeasure N A"  | 
|
280  | 
using A by (simp add: emeasure_nonneg positive_integral_cmult_indicator)  | 
|
281  | 
finally show ?thesis  | 
|
282  | 
by (simp add: ac_simps)  | 
|
| 40859 | 283  | 
qed  | 
284  | 
||
| 47694 | 285  | 
subsection {* Binary products of $\sigma$-finite emeasure spaces *}
 | 
| 40859 | 286  | 
|
| 47694 | 287  | 
locale pair_sigma_finite = M1: sigma_finite_measure M1 + M2: sigma_finite_measure M2  | 
288  | 
for M1 :: "'a measure" and M2 :: "'b measure"  | 
|
| 40859 | 289  | 
|
| 47694 | 290  | 
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
 | 
291  | 
"Q \<in> sets (M1 \<Otimes>\<^sub>M M2) \<Longrightarrow> (\<lambda>x. emeasure M2 (Pair x -` Q)) \<in> borel_measurable M1"  | 
| 49776 | 292  | 
using M2.measurable_emeasure_Pair .  | 
| 40859 | 293  | 
|
| 47694 | 294  | 
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
 | 
295  | 
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 | 296  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
297  | 
have "(\<lambda>(x, y). (y, x)) -` Q \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"  | 
| 47694 | 298  | 
using Q measurable_pair_swap' by (auto intro: measurable_sets)  | 
| 49776 | 299  | 
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
 | 
300  | 
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
 | 
301  | 
using Q[THEN sets.sets_into_space] by (auto simp: space_pair_measure)  | 
| 47694 | 302  | 
ultimately show ?thesis by simp  | 
| 
39088
 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 
hoelzl 
parents: 
39082 
diff
changeset
 | 
303  | 
qed  | 
| 
 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 
hoelzl 
parents: 
39082 
diff
changeset
 | 
304  | 
|
| 
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
 | 
305  | 
lemma (in pair_sigma_finite) sigma_finite_up_in_pair_measure_generator:  | 
| 47694 | 306  | 
  defines "E \<equiv> {A \<times> B | A B. A \<in> sets M1 \<and> B \<in> sets M2}"
 | 
307  | 
  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
 | 
308  | 
(\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)"  | 
| 40859 | 309  | 
proof -  | 
| 47694 | 310  | 
from M1.sigma_finite_incseq guess F1 . note F1 = this  | 
311  | 
from M2.sigma_finite_incseq guess F2 . note F2 = this  | 
|
312  | 
from F1 F2 have space: "space M1 = (\<Union>i. F1 i)" "space M2 = (\<Union>i. F2 i)" by auto  | 
|
| 40859 | 313  | 
let ?F = "\<lambda>i. F1 i \<times> F2 i"  | 
| 47694 | 314  | 
show ?thesis  | 
| 40859 | 315  | 
proof (intro exI[of _ ?F] conjI allI)  | 
| 47694 | 316  | 
show "range ?F \<subseteq> E" using F1 F2 by (auto simp: E_def) (metis range_subsetD)  | 
| 40859 | 317  | 
next  | 
318  | 
have "space M1 \<times> space M2 \<subseteq> (\<Union>i. ?F i)"  | 
|
319  | 
proof (intro subsetI)  | 
|
320  | 
fix x assume "x \<in> space M1 \<times> space M2"  | 
|
321  | 
then obtain i j where "fst x \<in> F1 i" "snd x \<in> F2 j"  | 
|
322  | 
by (auto simp: space)  | 
|
323  | 
then have "fst x \<in> F1 (max i j)" "snd x \<in> F2 (max j i)"  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
324  | 
using `incseq F1` `incseq F2` unfolding incseq_def  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
325  | 
by (force split: split_max)+  | 
| 40859 | 326  | 
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
 | 
327  | 
by (intro SigmaI) (auto simp add: max.commute)  | 
| 40859 | 328  | 
then show "x \<in> (\<Union>i. ?F i)" by auto  | 
329  | 
qed  | 
|
| 47694 | 330  | 
then show "(\<Union>i. ?F i) = space M1 \<times> space M2"  | 
331  | 
using space by (auto simp: space)  | 
|
| 40859 | 332  | 
next  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
333  | 
fix i show "incseq (\<lambda>i. F1 i \<times> F2 i)"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
334  | 
using `incseq F1` `incseq F2` unfolding incseq_Suc_iff by auto  | 
| 40859 | 335  | 
next  | 
336  | 
fix i  | 
|
337  | 
from F1 F2 have "F1 i \<in> sets M1" "F2 i \<in> sets M2" by auto  | 
|
| 47694 | 338  | 
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
 | 
339  | 
show "emeasure (M1 \<Otimes>\<^sub>M M2) (F1 i \<times> F2 i) \<noteq> \<infinity>"  | 
| 47694 | 340  | 
by (auto simp add: emeasure_pair_measure_Times)  | 
341  | 
qed  | 
|
342  | 
qed  | 
|
343  | 
||
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
344  | 
sublocale pair_sigma_finite \<subseteq> P: sigma_finite_measure "M1 \<Otimes>\<^sub>M M2"  | 
| 47694 | 345  | 
proof  | 
346  | 
  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
 | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
347  | 
  show "\<exists>F::nat \<Rightarrow> ('a \<times> 'b) set. range F \<subseteq> sets (M1 \<Otimes>\<^sub>M M2) \<and> (\<Union>i. F i) = space (M1 \<Otimes>\<^sub>M M2) \<and> (\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>)"
 | 
| 47694 | 348  | 
proof (rule exI[of _ F], intro conjI)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
349  | 
show "range F \<subseteq> sets (M1 \<Otimes>\<^sub>M M2)" using F by (auto simp: pair_measure_def)  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
350  | 
show "(\<Union>i. F i) = space (M1 \<Otimes>\<^sub>M M2)"  | 
| 47694 | 351  | 
using F 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
 | 
352  | 
show "\<forall>i. emeasure (M1 \<Otimes>\<^sub>M M2) (F i) \<noteq> \<infinity>" using F by auto  | 
| 40859 | 353  | 
qed  | 
354  | 
qed  | 
|
355  | 
||
| 47694 | 356  | 
lemma sigma_finite_pair_measure:  | 
357  | 
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
 | 
358  | 
shows "sigma_finite_measure (A \<Otimes>\<^sub>M B)"  | 
| 47694 | 359  | 
proof -  | 
360  | 
interpret A: sigma_finite_measure A by fact  | 
|
361  | 
interpret B: sigma_finite_measure B by fact  | 
|
362  | 
interpret AB: pair_sigma_finite A B ..  | 
|
363  | 
show ?thesis ..  | 
|
| 40859 | 364  | 
qed  | 
| 
39088
 
ca17017c10e6
Measurable on product space is equiv. to measurable components
 
hoelzl 
parents: 
39082 
diff
changeset
 | 
365  | 
|
| 47694 | 366  | 
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
 | 
367  | 
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
 | 
368  | 
shows "(\<lambda>(x, y). (y, x)) -` A \<inter> space (M2 \<Otimes>\<^sub>M M1) \<in> sets (M2 \<Otimes>\<^sub>M M1)"  | 
| 47694 | 369  | 
using measurable_pair_swap' assms by (rule measurable_sets)  | 
| 41661 | 370  | 
|
| 47694 | 371  | 
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
 | 
372  | 
"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 | 373  | 
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
 | 
374  | 
  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
 | 
| 47694 | 375  | 
  let ?E = "{a \<times> b |a b. a \<in> sets M1 \<and> b \<in> sets M2}"
 | 
376  | 
show ?thesis  | 
|
377  | 
proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])  | 
|
378  | 
show "?E \<subseteq> Pow (space ?P)"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
379  | 
using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure)  | 
| 47694 | 380  | 
show "sets ?P = sigma_sets (space ?P) ?E"  | 
381  | 
by (simp add: sets_pair_measure space_pair_measure)  | 
|
382  | 
then show "sets ?D = sigma_sets (space ?P) ?E"  | 
|
383  | 
by simp  | 
|
384  | 
next  | 
|
| 
49784
 
5e5b2da42a69
remove incseq assumption from measure_eqI_generator_eq
 
hoelzl 
parents: 
49776 
diff
changeset
 | 
385  | 
show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"  | 
| 47694 | 386  | 
using F by (auto simp: space_pair_measure)  | 
387  | 
next  | 
|
388  | 
fix X assume "X \<in> ?E"  | 
|
389  | 
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
 | 
390  | 
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
 | 
391  | 
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
 | 
392  | 
with A B show "emeasure (M1 \<Otimes>\<^sub>M M2) X = emeasure ?D X"  | 
| 49776 | 393  | 
by (simp add: M2.emeasure_pair_measure_Times M1.emeasure_pair_measure_Times emeasure_distr  | 
| 47694 | 394  | 
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
 | 
395  | 
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
 | 
396  | 
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
 | 
397  | 
|
| 47694 | 398  | 
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
 | 
399  | 
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
 | 
400  | 
shows "emeasure (M1 \<Otimes>\<^sub>M M2) A = (\<integral>\<^sup>+y. emeasure M1 ((\<lambda>x. (x, y)) -` A) \<partial>M2)"  | 
| 47694 | 401  | 
(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
 | 
402  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
403  | 
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
 | 
404  | 
using sets.sets_into_space[OF A] by (auto simp: space_pair_measure)  | 
| 47694 | 405  | 
show ?thesis using A  | 
406  | 
by (subst distr_pair_swap)  | 
|
407  | 
(simp_all del: vimage_Int add: measurable_sets[OF measurable_pair_swap']  | 
|
| 49776 | 408  | 
M1.emeasure_pair_measure_alt emeasure_distr[OF measurable_pair_swap' A])  | 
409  | 
qed  | 
|
410  | 
||
411  | 
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
 | 
412  | 
assumes "AE x in (M1 \<Otimes>\<^sub>M M2). Q x"  | 
| 49776 | 413  | 
shows "AE x in M1. (AE y in M2. Q (x, y))"  | 
414  | 
proof -  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
415  | 
  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 | 416  | 
using assms unfolding eventually_ae_filter by auto  | 
417  | 
show ?thesis  | 
|
418  | 
proof (rule AE_I)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
419  | 
from N measurable_emeasure_Pair1[OF `N \<in> sets (M1 \<Otimes>\<^sub>M M2)`]  | 
| 49776 | 420  | 
    show "emeasure M1 {x\<in>space M1. emeasure M2 (Pair x -` N) \<noteq> 0} = 0"
 | 
421  | 
by (auto simp: M2.emeasure_pair_measure_alt positive_integral_0_iff emeasure_nonneg)  | 
|
422  | 
    show "{x \<in> space M1. emeasure M2 (Pair x -` N) \<noteq> 0} \<in> sets M1"
 | 
|
423  | 
by (intro borel_measurable_ereal_neq_const measurable_emeasure_Pair1 N)  | 
|
424  | 
    { fix x assume "x \<in> space M1" "emeasure M2 (Pair x -` N) = 0"
 | 
|
425  | 
have "AE y in M2. Q (x, y)"  | 
|
426  | 
proof (rule AE_I)  | 
|
427  | 
show "emeasure M2 (Pair x -` N) = 0" by fact  | 
|
428  | 
show "Pair x -` N \<in> sets M2" using N(1) by (rule sets_Pair1)  | 
|
429  | 
        show "{y \<in> space M2. \<not> Q (x, y)} \<subseteq> Pair x -` N"
 | 
|
430  | 
using N `x \<in> space M1` unfolding space_pair_measure by auto  | 
|
431  | 
qed }  | 
|
432  | 
    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}"
 | 
|
433  | 
by auto  | 
|
434  | 
qed  | 
|
435  | 
qed  | 
|
436  | 
||
437  | 
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
 | 
438  | 
  assumes "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P x} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
| 49776 | 439  | 
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
 | 
440  | 
shows "AE x in M1 \<Otimes>\<^sub>M M2. P x"  | 
| 49776 | 441  | 
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
 | 
442  | 
  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
 | 
443  | 
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
 | 
444  | 
  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
 | 
445  | 
      (\<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 | 446  | 
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
 | 
447  | 
also have "\<dots> = (\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. 0 \<partial>M2 \<partial>M1)"  | 
| 49776 | 448  | 
using ae  | 
449  | 
apply (safe intro!: positive_integral_cong_AE)  | 
|
450  | 
apply (intro AE_I2)  | 
|
451  | 
apply (safe intro!: positive_integral_cong_AE)  | 
|
452  | 
apply auto  | 
|
453  | 
done  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
454  | 
  finally show "emeasure (M1 \<Otimes>\<^sub>M M2) {x \<in> space (M1 \<Otimes>\<^sub>M M2). \<not> P x} = 0" by simp
 | 
| 49776 | 455  | 
qed  | 
456  | 
||
457  | 
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
 | 
458  | 
  "{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
 | 
459  | 
(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 | 460  | 
using AE_pair[of "\<lambda>x. P (fst x) (snd x)"] AE_pair_measure[of "\<lambda>x. P (fst x) (snd x)"] by auto  | 
461  | 
||
462  | 
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
 | 
463  | 
  assumes P: "{x\<in>space (M1 \<Otimes>\<^sub>M M2). P (fst x) (snd x)} \<in> sets (M1 \<Otimes>\<^sub>M M2)"
 | 
| 49776 | 464  | 
shows "(AE x in M1. AE y in M2. P x y) \<longleftrightarrow> (AE y in M2. AE x in M1. P x y)"  | 
465  | 
proof -  | 
|
466  | 
interpret Q: pair_sigma_finite M2 M1 ..  | 
|
467  | 
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"  | 
|
468  | 
by auto  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
469  | 
  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
 | 
470  | 
    (\<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 | 471  | 
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
 | 
472  | 
also have "\<dots> \<in> sets (M2 \<Otimes>\<^sub>M M1)"  | 
| 49776 | 473  | 
by (intro sets_pair_swap P)  | 
474  | 
finally show ?thesis  | 
|
475  | 
apply (subst AE_pair_iff[OF P])  | 
|
476  | 
apply (subst distr_pair_swap)  | 
|
477  | 
apply (subst AE_distr_iff[OF measurable_pair_swap' P])  | 
|
478  | 
apply (subst Q.AE_pair_iff)  | 
|
479  | 
apply simp_all  | 
|
480  | 
done  | 
|
| 40859 | 481  | 
qed  | 
482  | 
||
483  | 
section "Fubinis theorem"  | 
|
484  | 
||
| 49800 | 485  | 
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
 | 
486  | 
"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 | 487  | 
by simp  | 
| 49800 | 488  | 
|
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
489  | 
lemma (in sigma_finite_measure) borel_measurable_positive_integral_fst':  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
490  | 
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
 | 
491  | 
shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"  | 
| 49800 | 492  | 
using f proof induct  | 
493  | 
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
 | 
494  | 
then have "\<And>w x. w \<in> space M1 \<Longrightarrow> x \<in> space M \<Longrightarrow> u (w, x) = v (w, x)"  | 
| 49800 | 495  | 
by (auto simp: space_pair_measure)  | 
496  | 
show ?case  | 
|
497  | 
apply (subst measurable_cong)  | 
|
498  | 
apply (rule positive_integral_cong)  | 
|
499  | 
apply fact+  | 
|
500  | 
done  | 
|
501  | 
next  | 
|
502  | 
case (set Q)  | 
|
503  | 
have [simp]: "\<And>x y. indicator Q (x, y) = indicator (Pair x -` Q) y"  | 
|
504  | 
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
 | 
505  | 
have "\<And>x. x \<in> space M1 \<Longrightarrow> emeasure M (Pair x -` Q) = \<integral>\<^sup>+ y. indicator Q (x, y) \<partial>M"  | 
| 49800 | 506  | 
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
 | 
507  | 
from this measurable_emeasure_Pair[OF set] show ?case  | 
| 49800 | 508  | 
by (rule measurable_cong[THEN iffD1])  | 
509  | 
qed (simp_all add: positive_integral_add positive_integral_cmult measurable_compose_Pair1  | 
|
510  | 
positive_integral_monotone_convergence_SUP incseq_def le_fun_def  | 
|
511  | 
cong: measurable_cong)  | 
|
512  | 
||
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
513  | 
lemma (in sigma_finite_measure) positive_integral_fst:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
514  | 
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
 | 
515  | 
shows "(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f (x, y) \<partial>M \<partial>M1) = integral\<^sup>P (M1 \<Otimes>\<^sub>M M) f" (is "?I f = _")  | 
| 49800 | 516  | 
using f proof induct  | 
517  | 
case (cong u v)  | 
|
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
518  | 
then have "?I u = ?I v"  | 
| 49800 | 519  | 
by (intro positive_integral_cong) (auto simp: space_pair_measure)  | 
| 
53374
 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 
wenzelm 
parents: 
53015 
diff
changeset
 | 
520  | 
with cong show ?case  | 
| 49800 | 521  | 
by (simp cong: positive_integral_cong)  | 
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
522  | 
qed (simp_all add: emeasure_pair_measure positive_integral_cmult positive_integral_add  | 
| 49800 | 523  | 
positive_integral_monotone_convergence_SUP  | 
524  | 
measurable_compose_Pair1 positive_integral_positive  | 
|
| 
49825
 
bb5db3d1d6dd
cleanup borel_measurable_positive_integral_(fst|snd)
 
hoelzl 
parents: 
49800 
diff
changeset
 | 
525  | 
borel_measurable_positive_integral_fst' positive_integral_mono incseq_def le_fun_def  | 
| 49800 | 526  | 
cong: positive_integral_cong)  | 
| 40859 | 527  | 
|
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
528  | 
lemma (in sigma_finite_measure) positive_integral_fst_measurable:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
529  | 
assumes f: "f \<in> borel_measurable (M1 \<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
 | 
530  | 
shows "(\<lambda>x. \<integral>\<^sup>+ y. f (x, y) \<partial>M) \<in> borel_measurable M1"  | 
| 40859 | 531  | 
(is "?C f \<in> borel_measurable M1")  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
532  | 
and "(\<integral>\<^sup>+ x. (\<integral>\<^sup>+ y. f (x, y) \<partial>M) \<partial>M1) = integral\<^sup>P (M1 \<Otimes>\<^sub>M M) f"  | 
| 49800 | 533  | 
using f  | 
| 
49825
 
bb5db3d1d6dd
cleanup borel_measurable_positive_integral_(fst|snd)
 
hoelzl 
parents: 
49800 
diff
changeset
 | 
534  | 
borel_measurable_positive_integral_fst'[of "\<lambda>x. max 0 (f x)"]  | 
| 49800 | 535  | 
positive_integral_fst[of "\<lambda>x. max 0 (f x)"]  | 
536  | 
unfolding positive_integral_max_0 by auto  | 
|
| 40859 | 537  | 
|
| 50003 | 538  | 
lemma (in sigma_finite_measure) borel_measurable_positive_integral[measurable (raw)]:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
539  | 
"split f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral>\<^sup>+ y. f x y \<partial>M) \<in> borel_measurable N"  | 
| 50003 | 540  | 
using positive_integral_fst_measurable(1)[of "split f" N] by simp  | 
541  | 
||
542  | 
lemma (in sigma_finite_measure) borel_measurable_lebesgue_integral[measurable (raw)]:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
543  | 
"split f \<in> borel_measurable (N \<Otimes>\<^sub>M M) \<Longrightarrow> (\<lambda>x. \<integral> y. f x y \<partial>M) \<in> borel_measurable N"  | 
| 50003 | 544  | 
by (simp add: lebesgue_integral_def)  | 
| 
49825
 
bb5db3d1d6dd
cleanup borel_measurable_positive_integral_(fst|snd)
 
hoelzl 
parents: 
49800 
diff
changeset
 | 
545  | 
|
| 47694 | 546  | 
lemma (in pair_sigma_finite) positive_integral_snd_measurable:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
547  | 
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
 | 
548  | 
shows "(\<integral>\<^sup>+ y. (\<integral>\<^sup>+ x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) f"  | 
| 41661 | 549  | 
proof -  | 
| 47694 | 550  | 
note measurable_pair_swap[OF f]  | 
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
551  | 
from M1.positive_integral_fst_measurable[OF this]  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
552  | 
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 | 553  | 
by simp  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
554  | 
also have "(\<integral>\<^sup>+ (x, y). f (y, x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) f"  | 
| 47694 | 555  | 
by (subst distr_pair_swap)  | 
556  | 
(auto simp: positive_integral_distr[OF measurable_pair_swap' f] intro!: positive_integral_cong)  | 
|
| 40859 | 557  | 
finally show ?thesis .  | 
558  | 
qed  | 
|
559  | 
||
560  | 
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
 | 
561  | 
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
 | 
562  | 
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)"  | 
| 40859 | 563  | 
unfolding positive_integral_snd_measurable[OF assms]  | 
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
564  | 
unfolding M2.positive_integral_fst_measurable[OF assms] ..  | 
| 40859 | 565  | 
|
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
566  | 
lemma (in pair_sigma_finite) integrable_product_swap:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
567  | 
assumes "integrable (M1 \<Otimes>\<^sub>M M2) f"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
568  | 
shows "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x))"  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
569  | 
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
 | 
570  | 
interpret Q: pair_sigma_finite M2 M1 by default  | 
| 41661 | 571  | 
have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)  | 
572  | 
show ?thesis unfolding *  | 
|
| 47694 | 573  | 
by (rule integrable_distr[OF measurable_pair_swap'])  | 
574  | 
(simp add: distr_pair_swap[symmetric] assms)  | 
|
| 41661 | 575  | 
qed  | 
576  | 
||
577  | 
lemma (in pair_sigma_finite) integrable_product_swap_iff:  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
578  | 
"integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x,y). f (y,x)) \<longleftrightarrow> integrable (M1 \<Otimes>\<^sub>M M2) f"  | 
| 41661 | 579  | 
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
 | 
580  | 
interpret Q: pair_sigma_finite M2 M1 by default  | 
| 41661 | 581  | 
from Q.integrable_product_swap[of "\<lambda>(x,y). f (y,x)"] integrable_product_swap[of f]  | 
582  | 
show ?thesis by auto  | 
|
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
583  | 
qed  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
584  | 
|
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
585  | 
lemma (in pair_sigma_finite) integral_product_swap:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
586  | 
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
 | 
587  | 
shows "(\<integral>(x,y). f (y,x) \<partial>(M2 \<Otimes>\<^sub>M M1)) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f"  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
588  | 
proof -  | 
| 41661 | 589  | 
have *: "(\<lambda>(x,y). f (y,x)) = (\<lambda>x. f (case x of (x,y)\<Rightarrow>(y,x)))" by (auto simp: fun_eq_iff)  | 
| 47694 | 590  | 
show ?thesis unfolding *  | 
591  | 
by (simp add: integral_distr[symmetric, OF measurable_pair_swap' f] distr_pair_swap[symmetric])  | 
|
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
592  | 
qed  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
593  | 
|
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
594  | 
lemma (in pair_sigma_finite) integrable_fst_measurable:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
595  | 
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f"  | 
| 47694 | 596  | 
shows "AE x in M1. integrable M2 (\<lambda> y. f (x, y))" (is "?AE")  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
597  | 
and "(\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f" (is "?INT")  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
598  | 
proof -  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
599  | 
have f_borel: "f \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)"  | 
| 47694 | 600  | 
using f by auto  | 
| 46731 | 601  | 
let ?pf = "\<lambda>x. ereal (f x)" and ?nf = "\<lambda>x. ereal (- f x)"  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
602  | 
have  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
603  | 
borel: "?nf \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)""?pf \<in> borel_measurable (M1 \<Otimes>\<^sub>M M2)" and  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
604  | 
int: "integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) ?nf \<noteq> \<infinity>" "integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) ?pf \<noteq> \<infinity>"  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
605  | 
using assms by auto  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
606  | 
have "(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
607  | 
"(\<integral>\<^sup>+x. (\<integral>\<^sup>+y. ereal (- f (x, y)) \<partial>M2) \<partial>M1) \<noteq> \<infinity>"  | 
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
608  | 
using borel[THEN M2.positive_integral_fst_measurable(1)] int  | 
| 
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
609  | 
unfolding borel[THEN M2.positive_integral_fst_measurable(2)] by simp_all  | 
| 
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
610  | 
with borel[THEN M2.positive_integral_fst_measurable(1)]  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
611  | 
have AE_pos: "AE x in M1. (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
612  | 
"AE x in M1. (\<integral>\<^sup>+y. ereal (- f (x, y)) \<partial>M2) \<noteq> \<infinity>"  | 
| 47694 | 613  | 
by (auto intro!: positive_integral_PInf_AE )  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
614  | 
then have AE: "AE x in M1. \<bar>\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
615  | 
"AE x in M1. \<bar>\<integral>\<^sup>+y. ereal (- f (x, y)) \<partial>M2\<bar> \<noteq> \<infinity>"  | 
| 47694 | 616  | 
by (auto simp: positive_integral_positive)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
617  | 
from AE_pos show ?AE using assms  | 
| 47694 | 618  | 
by (simp add: measurable_Pair2[OF f_borel] integrable_def)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
619  | 
  { fix f have "(\<integral>\<^sup>+ x. - \<integral>\<^sup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = (\<integral>\<^sup>+x. 0 \<partial>M1)"
 | 
| 47694 | 620  | 
using positive_integral_positive  | 
621  | 
by (intro positive_integral_cong_pos) (auto simp: ereal_uminus_le_reorder)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
622  | 
then have "(\<integral>\<^sup>+ x. - \<integral>\<^sup>+ y. ereal (f x y) \<partial>M2 \<partial>M1) = 0" by simp }  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
623  | 
note this[simp]  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
624  | 
  { fix f assume borel: "(\<lambda>x. ereal (f x)) \<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
 | 
625  | 
and int: "integral\<^sup>P (M1 \<Otimes>\<^sub>M M2) (\<lambda>x. ereal (f x)) \<noteq> \<infinity>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
626  | 
and AE: "AE x in M1. (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2) \<noteq> \<infinity>"  | 
| 
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
627  | 
have "integrable M1 (\<lambda>x. real (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2))" (is "integrable M1 ?f")  | 
| 41705 | 628  | 
proof (intro integrable_def[THEN iffD2] conjI)  | 
629  | 
show "?f \<in> borel_measurable M1"  | 
|
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
630  | 
using borel by (auto intro!: M2.positive_integral_fst_measurable)  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
631  | 
have "(\<integral>\<^sup>+x. ereal (?f x) \<partial>M1) = (\<integral>\<^sup>+x. (\<integral>\<^sup>+y. ereal (f (x, y)) \<partial>M2) \<partial>M1)"  | 
| 47694 | 632  | 
using AE positive_integral_positive[of M2]  | 
633  | 
by (auto intro!: positive_integral_cong_AE simp: ereal_real)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
634  | 
then show "(\<integral>\<^sup>+x. ereal (?f x) \<partial>M1) \<noteq> \<infinity>"  | 
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
635  | 
using M2.positive_integral_fst_measurable[OF borel] int by simp  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
636  | 
have "(\<integral>\<^sup>+x. ereal (- ?f x) \<partial>M1) = (\<integral>\<^sup>+x. 0 \<partial>M1)"  | 
| 47694 | 637  | 
by (intro positive_integral_cong_pos)  | 
638  | 
(simp add: positive_integral_positive real_of_ereal_pos)  | 
|
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
639  | 
then show "(\<integral>\<^sup>+x. ereal (- ?f x) \<partial>M1) \<noteq> \<infinity>" by simp  | 
| 41705 | 640  | 
qed }  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
641  | 
with this[OF borel(1) int(1) AE_pos(2)] this[OF borel(2) int(2) AE_pos(1)]  | 
| 41705 | 642  | 
show ?INT  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
643  | 
unfolding lebesgue_integral_def[of "M1 \<Otimes>\<^sub>M M2"] lebesgue_integral_def[of M2]  | 
| 
49999
 
dfb63b9b8908
for the product measure it is enough if only one measure is sigma-finite
 
hoelzl 
parents: 
49825 
diff
changeset
 | 
644  | 
borel[THEN M2.positive_integral_fst_measurable(2), symmetric]  | 
| 47694 | 645  | 
using AE[THEN integral_real]  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
646  | 
by simp  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
647  | 
qed  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
648  | 
|
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
649  | 
lemma (in pair_sigma_finite) integrable_snd_measurable:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
650  | 
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f"  | 
| 47694 | 651  | 
shows "AE y in M2. integrable M1 (\<lambda>x. f (x, y))" (is "?AE")  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
652  | 
and "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = integral\<^sup>L (M1 \<Otimes>\<^sub>M M2) f" (is "?INT")  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
653  | 
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
 | 
654  | 
interpret Q: pair_sigma_finite M2 M1 by default  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
655  | 
have Q_int: "integrable (M2 \<Otimes>\<^sub>M M1) (\<lambda>(x, y). f (y, x))"  | 
| 41661 | 656  | 
using f unfolding integrable_product_swap_iff .  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
657  | 
show ?INT  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
658  | 
using Q.integrable_fst_measurable(2)[OF Q_int]  | 
| 47694 | 659  | 
using integral_product_swap[of f] f by auto  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
660  | 
show ?AE  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
661  | 
using Q.integrable_fst_measurable(1)[OF Q_int]  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
662  | 
by simp  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
663  | 
qed  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
664  | 
|
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
665  | 
lemma (in pair_sigma_finite) Fubini_integral:  | 
| 
53015
 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 
wenzelm 
parents: 
50244 
diff
changeset
 | 
666  | 
assumes f: "integrable (M1 \<Otimes>\<^sub>M M2) f"  | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41661 
diff
changeset
 | 
667  | 
shows "(\<integral>y. (\<integral>x. f (x, y) \<partial>M1) \<partial>M2) = (\<integral>x. (\<integral>y. f (x, y) \<partial>M2) \<partial>M1)"  | 
| 
41026
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
668  | 
unfolding integrable_snd_measurable[OF assms]  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
669  | 
unfolding integrable_fst_measurable[OF assms] ..  | 
| 
 
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
 
hoelzl 
parents: 
41023 
diff
changeset
 | 
670  | 
|
| 47694 | 671  | 
section {* Products on counting spaces, densities and distributions *}
 | 
| 40859 | 672  | 
|
| 
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
 | 
673  | 
lemma sigma_sets_pair_measure_generator_finite:  | 
| 38656 | 674  | 
assumes "finite A" and "finite B"  | 
| 47694 | 675  | 
  shows "sigma_sets (A \<times> B) { a \<times> b | a b. a \<subseteq> A \<and> b \<subseteq> B} = Pow (A \<times> B)"
 | 
| 40859 | 676  | 
(is "sigma_sets ?prod ?sets = _")  | 
| 38656 | 677  | 
proof safe  | 
678  | 
have fin: "finite (A \<times> B)" using assms by (rule finite_cartesian_product)  | 
|
679  | 
fix x assume subset: "x \<subseteq> A \<times> B"  | 
|
680  | 
hence "finite x" using fin by (rule finite_subset)  | 
|
| 40859 | 681  | 
from this subset show "x \<in> sigma_sets ?prod ?sets"  | 
| 38656 | 682  | 
proof (induct x)  | 
683  | 
case empty show ?case by (rule sigma_sets.Empty)  | 
|
684  | 
next  | 
|
685  | 
case (insert a x)  | 
|
| 47694 | 686  | 
    hence "{a} \<in> sigma_sets ?prod ?sets" by auto
 | 
| 38656 | 687  | 
moreover have "x \<in> sigma_sets ?prod ?sets" using insert by auto  | 
688  | 
ultimately show ?case unfolding insert_is_Un[of a x] by (rule sigma_sets_Un)  | 
|
689  | 
qed  | 
|
690  | 
next  | 
|
691  | 
fix x a b  | 
|
| 40859 | 692  | 
assume "x \<in> sigma_sets ?prod ?sets" and "(a, b) \<in> x"  | 
| 38656 | 693  | 
from sigma_sets_into_sp[OF _ this(1)] this(2)  | 
| 40859 | 694  | 
show "a \<in> A" and "b \<in> B" by auto  | 
| 35833 | 695  | 
qed  | 
696  | 
||
| 47694 | 697  | 
lemma pair_measure_count_space:  | 
698  | 
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
 | 
699  | 
shows "count_space A \<Otimes>\<^sub>M count_space B = count_space (A \<times> B)" (is "?P = ?C")  | 
| 47694 | 700  | 
proof (rule measure_eqI)  | 
701  | 
interpret A: finite_measure "count_space A" by (rule finite_measure_count_space) fact  | 
|
702  | 
interpret B: finite_measure "count_space B" by (rule finite_measure_count_space) fact  | 
|
703  | 
interpret P: pair_sigma_finite "count_space A" "count_space B" by default  | 
|
704  | 
show eq: "sets ?P = sets ?C"  | 
|
705  | 
by (simp add: sets_pair_measure sigma_sets_pair_measure_generator_finite A B)  | 
|
706  | 
fix X assume X: "X \<in> sets ?P"  | 
|
707  | 
with eq have X_subset: "X \<subseteq> A \<times> B" by simp  | 
|
708  | 
with A B have fin_Pair: "\<And>x. finite (Pair x -` X)"  | 
|
709  | 
by (intro finite_subset[OF _ B]) auto  | 
|
710  | 
have fin_X: "finite X" using X_subset by (rule finite_subset) (auto simp: A B)  | 
|
711  | 
show "emeasure ?P X = emeasure ?C X"  | 
|
| 49776 | 712  | 
apply (subst B.emeasure_pair_measure_alt[OF X])  | 
| 47694 | 713  | 
apply (subst emeasure_count_space)  | 
714  | 
using X_subset apply auto []  | 
|
715  | 
apply (simp add: fin_Pair emeasure_count_space X_subset fin_X)  | 
|
716  | 
apply (subst positive_integral_count_space)  | 
|
717  | 
using A apply simp  | 
|
718  | 
apply (simp del: real_of_nat_setsum add: real_of_nat_setsum[symmetric])  | 
|
719  | 
apply (subst card_gt_0_iff)  | 
|
720  | 
apply (simp add: fin_Pair)  | 
|
721  | 
apply (subst card_SigmaI[symmetric])  | 
|
722  | 
using A apply simp  | 
|
723  | 
using fin_Pair apply simp  | 
|
724  | 
using X_subset apply (auto intro!: arg_cong[where f=card])  | 
|
725  | 
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
 | 
726  | 
qed  | 
| 35833 | 727  | 
|
| 47694 | 728  | 
lemma pair_measure_density:  | 
729  | 
assumes f: "f \<in> borel_measurable M1" "AE x in M1. 0 \<le> f x"  | 
|
730  | 
assumes g: "g \<in> borel_measurable M2" "AE x in M2. 0 \<le> g x"  | 
|
| 50003 | 731  | 
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
 | 
732  | 
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 | 733  | 
proof (rule measure_eqI)  | 
734  | 
interpret M2: sigma_finite_measure M2 by fact  | 
|
735  | 
interpret D2: sigma_finite_measure "density M2 g" by fact  | 
|
736  | 
||
737  | 
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
 | 
738  | 
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
 | 
739  | 
(\<integral>\<^sup>+ x. \<integral>\<^sup>+ y. f x * g y * indicator A (x, y) \<partial>M2 \<partial>M1)"  | 
| 50003 | 740  | 
by (intro positive_integral_cong_AE)  | 
741  | 
(auto simp add: positive_integral_cmult[symmetric] ac_simps)  | 
|
742  | 
with A f g show "emeasure ?L A = emeasure ?R A"  | 
|
743  | 
by (simp add: D2.emeasure_pair_measure emeasure_density positive_integral_density  | 
|
744  | 
M2.positive_integral_fst_measurable(2)[symmetric]  | 
|
745  | 
cong: positive_integral_cong)  | 
|
| 47694 | 746  | 
qed simp  | 
747  | 
||
748  | 
lemma sigma_finite_measure_distr:  | 
|
749  | 
assumes "sigma_finite_measure (distr M N f)" and f: "f \<in> measurable M N"  | 
|
750  | 
shows "sigma_finite_measure M"  | 
|
| 40859 | 751  | 
proof -  | 
| 47694 | 752  | 
interpret sigma_finite_measure "distr M N f" by fact  | 
753  | 
from sigma_finite_disjoint guess A . note A = this  | 
|
754  | 
show ?thesis  | 
|
755  | 
proof (unfold_locales, intro conjI exI allI)  | 
|
756  | 
show "range (\<lambda>i. f -` A i \<inter> space M) \<subseteq> sets M"  | 
|
| 50003 | 757  | 
using A f by auto  | 
| 47694 | 758  | 
show "(\<Union>i. f -` A i \<inter> space M) = space M"  | 
759  | 
using A(1) A(2)[symmetric] f by (auto simp: measurable_def Pi_def)  | 
|
760  | 
fix i show "emeasure M (f -` A i \<inter> space M) \<noteq> \<infinity>"  | 
|
761  | 
using f A(1,2) A(3)[of i] by (simp add: emeasure_distr subset_eq)  | 
|
762  | 
qed  | 
|
| 38656 | 763  | 
qed  | 
764  | 
||
| 47694 | 765  | 
lemma pair_measure_distr:  | 
766  | 
assumes f: "f \<in> measurable M S" and g: "g \<in> measurable N T"  | 
|
| 50003 | 767  | 
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
 | 
768  | 
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 | 769  | 
proof (rule measure_eqI)  | 
770  | 
interpret T: sigma_finite_measure "distr N T g" by fact  | 
|
771  | 
interpret N: sigma_finite_measure N by (rule sigma_finite_measure_distr) fact+  | 
|
| 50003 | 772  | 
|
| 47694 | 773  | 
fix A assume A: "A \<in> sets ?P"  | 
| 50003 | 774  | 
with f g show "emeasure ?P A = emeasure ?D A"  | 
775  | 
by (auto simp add: N.emeasure_pair_measure_alt space_pair_measure emeasure_distr  | 
|
776  | 
T.emeasure_pair_measure_alt positive_integral_distr  | 
|
777  | 
intro!: positive_integral_cong arg_cong[where f="emeasure N"])  | 
|
778  | 
qed simp  | 
|
| 39097 | 779  | 
|
| 50104 | 780  | 
lemma pair_measure_eqI:  | 
781  | 
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
 | 
782  | 
assumes sets: "sets (M1 \<Otimes>\<^sub>M M2) = sets M"  | 
| 50104 | 783  | 
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
 | 
784  | 
shows "M1 \<Otimes>\<^sub>M M2 = M"  | 
| 50104 | 785  | 
proof -  | 
786  | 
interpret M1: sigma_finite_measure M1 by fact  | 
|
787  | 
interpret M2: sigma_finite_measure M2 by fact  | 
|
788  | 
interpret pair_sigma_finite M1 M2 by default  | 
|
789  | 
  from sigma_finite_up_in_pair_measure_generator guess F :: "nat \<Rightarrow> ('a \<times> 'b) set" .. note F = this
 | 
|
790  | 
  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
 | 
791  | 
let ?P = "M1 \<Otimes>\<^sub>M M2"  | 
| 50104 | 792  | 
show ?thesis  | 
793  | 
proof (rule measure_eqI_generator_eq[OF Int_stable_pair_measure_generator[of M1 M2]])  | 
|
794  | 
show "?E \<subseteq> Pow (space ?P)"  | 
|
| 
50244
 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 
immler 
parents: 
50104 
diff
changeset
 | 
795  | 
using sets.space_closed[of M1] sets.space_closed[of M2] by (auto simp: space_pair_measure)  | 
| 50104 | 796  | 
show "sets ?P = sigma_sets (space ?P) ?E"  | 
797  | 
by (simp add: sets_pair_measure space_pair_measure)  | 
|
798  | 
then show "sets M = sigma_sets (space ?P) ?E"  | 
|
799  | 
using sets[symmetric] by simp  | 
|
800  | 
next  | 
|
801  | 
show "range F \<subseteq> ?E" "(\<Union>i. F i) = space ?P" "\<And>i. emeasure ?P (F i) \<noteq> \<infinity>"  | 
|
802  | 
using F by (auto simp: space_pair_measure)  | 
|
803  | 
next  | 
|
804  | 
fix X assume "X \<in> ?E"  | 
|
805  | 
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  | 
|
806  | 
then have "emeasure ?P X = emeasure M1 A * emeasure M2 B"  | 
|
807  | 
by (simp add: M2.emeasure_pair_measure_Times)  | 
|
808  | 
also have "\<dots> = emeasure M (A \<times> B)"  | 
|
809  | 
using A B emeasure by auto  | 
|
810  | 
finally show "emeasure ?P X = emeasure M X"  | 
|
811  | 
by simp  | 
|
812  | 
qed  | 
|
813  | 
qed  | 
|
814  | 
||
| 40859 | 815  | 
end  |