author | huffman |
Thu, 27 Oct 2011 07:46:57 +0200 | |
changeset 45270 | d5b5c9259afd |
parent 44928 | 7ef6505bde7f |
child 45777 | c36637603821 |
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/Finite_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 {*Finite product measures*} |
42067 | 6 |
|
42146
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
7 |
theory Finite_Product_Measure |
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
8 |
imports Binary_Product_Measure |
35833 | 9 |
begin |
10 |
||
42146
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
11 |
lemma Pi_iff: "f \<in> Pi I X \<longleftrightarrow> (\<forall>i\<in>I. f i \<in> X i)" |
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
hoelzl
parents:
42067
diff
changeset
|
12 |
unfolding Pi_def by auto |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
13 |
|
40859 | 14 |
abbreviation |
15 |
"Pi\<^isub>E A B \<equiv> Pi A B \<inter> extensional A" |
|
39094 | 16 |
|
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
|
17 |
syntax |
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
|
18 |
"_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3PIE _:_./ _)" 10) |
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
|
19 |
|
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
|
20 |
syntax (xsymbols) |
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
|
21 |
"_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^isub>E _\<in>_./ _)" 10) |
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
|
22 |
|
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
|
23 |
syntax (HTML output) |
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
|
24 |
"_PiE" :: "[pttrn, 'a set, 'b set] => ('a => 'b) set" ("(3\<Pi>\<^isub>E _\<in>_./ _)" 10) |
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
|
25 |
|
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
|
26 |
translations |
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 |
"PIE x:A. B" == "CONST Pi\<^isub>E A (%x. B)" |
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
|
28 |
|
40859 | 29 |
abbreviation |
30 |
funcset_extensional :: "['a set, 'b set] => ('a => 'b) set" |
|
31 |
(infixr "->\<^isub>E" 60) where |
|
32 |
"A ->\<^isub>E B \<equiv> Pi\<^isub>E A (%_. B)" |
|
33 |
||
34 |
notation (xsymbols) |
|
35 |
funcset_extensional (infixr "\<rightarrow>\<^isub>E" 60) |
|
36 |
||
37 |
lemma extensional_empty[simp]: "extensional {} = {\<lambda>x. undefined}" |
|
38 |
by safe (auto simp add: extensional_def fun_eq_iff) |
|
39 |
||
40 |
lemma extensional_insert[intro, simp]: |
|
41 |
assumes "a \<in> extensional (insert i I)" |
|
42 |
shows "a(i := b) \<in> extensional (insert i I)" |
|
43 |
using assms unfolding extensional_def by auto |
|
44 |
||
45 |
lemma extensional_Int[simp]: |
|
46 |
"extensional I \<inter> extensional I' = extensional (I \<inter> I')" |
|
47 |
unfolding extensional_def by auto |
|
38656 | 48 |
|
35833 | 49 |
definition |
40859 | 50 |
"merge I x J y = (\<lambda>i. if i \<in> I then x i else if i \<in> J then y i else undefined)" |
51 |
||
52 |
lemma merge_apply[simp]: |
|
53 |
"I \<inter> J = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I x J y i = x i" |
|
54 |
"I \<inter> J = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i" |
|
55 |
"J \<inter> I = {} \<Longrightarrow> i \<in> I \<Longrightarrow> merge I x J y i = x i" |
|
56 |
"J \<inter> I = {} \<Longrightarrow> i \<in> J \<Longrightarrow> merge I x J y i = y i" |
|
57 |
"i \<notin> I \<Longrightarrow> i \<notin> J \<Longrightarrow> merge I x J y i = undefined" |
|
58 |
unfolding merge_def by auto |
|
59 |
||
60 |
lemma merge_commute: |
|
61 |
"I \<inter> J = {} \<Longrightarrow> merge I x J y = merge J y I x" |
|
62 |
by (auto simp: merge_def intro!: ext) |
|
63 |
||
64 |
lemma Pi_cancel_merge_range[simp]: |
|
65 |
"I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A" |
|
66 |
"I \<inter> J = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A" |
|
67 |
"J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge I A J B) \<longleftrightarrow> x \<in> Pi I A" |
|
68 |
"J \<inter> I = {} \<Longrightarrow> x \<in> Pi I (merge J B I A) \<longleftrightarrow> x \<in> Pi I A" |
|
69 |
by (auto simp: Pi_def) |
|
70 |
||
71 |
lemma Pi_cancel_merge[simp]: |
|
72 |
"I \<inter> J = {} \<Longrightarrow> merge I x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B" |
|
73 |
"J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B" |
|
74 |
"I \<inter> J = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B" |
|
75 |
"J \<inter> I = {} \<Longrightarrow> merge I x J y \<in> Pi J B \<longleftrightarrow> y \<in> Pi J B" |
|
76 |
by (auto simp: Pi_def) |
|
77 |
||
78 |
lemma extensional_merge[simp]: "merge I x J y \<in> extensional (I \<union> J)" |
|
79 |
by (auto simp: extensional_def) |
|
80 |
||
81 |
lemma restrict_Pi_cancel: "restrict x I \<in> Pi I A \<longleftrightarrow> x \<in> Pi I A" |
|
82 |
by (auto simp: restrict_def Pi_def) |
|
83 |
||
84 |
lemma restrict_merge[simp]: |
|
85 |
"I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I" |
|
86 |
"I \<inter> J = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J" |
|
87 |
"J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) I = restrict x I" |
|
88 |
"J \<inter> I = {} \<Longrightarrow> restrict (merge I x J y) J = restrict y J" |
|
89 |
by (auto simp: restrict_def intro!: ext) |
|
90 |
||
91 |
lemma extensional_insert_undefined[intro, simp]: |
|
92 |
assumes "a \<in> extensional (insert i I)" |
|
93 |
shows "a(i := undefined) \<in> extensional I" |
|
94 |
using assms unfolding extensional_def by auto |
|
95 |
||
96 |
lemma extensional_insert_cancel[intro, simp]: |
|
97 |
assumes "a \<in> extensional I" |
|
98 |
shows "a \<in> extensional (insert i I)" |
|
99 |
using assms unfolding extensional_def by auto |
|
100 |
||
41095 | 101 |
lemma merge_singleton[simp]: "i \<notin> I \<Longrightarrow> merge I x {i} y = restrict (x(i := y i)) (insert i I)" |
102 |
unfolding merge_def by (auto simp: fun_eq_iff) |
|
103 |
||
104 |
lemma Pi_Int: "Pi I E \<inter> Pi I F = (\<Pi> i\<in>I. E i \<inter> F i)" |
|
105 |
by auto |
|
106 |
||
40859 | 107 |
lemma PiE_Int: "(Pi\<^isub>E I A) \<inter> (Pi\<^isub>E I B) = Pi\<^isub>E I (\<lambda>x. A x \<inter> B x)" |
108 |
by auto |
|
109 |
||
110 |
lemma Pi_cancel_fupd_range[simp]: "i \<notin> I \<Longrightarrow> x \<in> Pi I (B(i := b)) \<longleftrightarrow> x \<in> Pi I B" |
|
111 |
by (auto simp: Pi_def) |
|
112 |
||
113 |
lemma Pi_split_insert_domain[simp]: "x \<in> Pi (insert i I) X \<longleftrightarrow> x \<in> Pi I X \<and> x i \<in> X i" |
|
114 |
by (auto simp: Pi_def) |
|
39088
ca17017c10e6
Measurable on product space is equiv. to measurable components
hoelzl
parents:
39082
diff
changeset
|
115 |
|
40859 | 116 |
lemma Pi_split_domain[simp]: "x \<in> Pi (I \<union> J) X \<longleftrightarrow> x \<in> Pi I X \<and> x \<in> Pi J X" |
117 |
by (auto simp: Pi_def) |
|
118 |
||
119 |
lemma Pi_cancel_fupd[simp]: "i \<notin> I \<Longrightarrow> x(i := a) \<in> Pi I B \<longleftrightarrow> x \<in> Pi I B" |
|
120 |
by (auto simp: Pi_def) |
|
121 |
||
41095 | 122 |
lemma restrict_vimage: |
123 |
assumes "I \<inter> J = {}" |
|
124 |
shows "(\<lambda>x. (restrict x I, restrict x J)) -` (Pi\<^isub>E I E \<times> Pi\<^isub>E J F) = Pi (I \<union> J) (merge I E J F)" |
|
125 |
using assms by (auto simp: restrict_Pi_cancel) |
|
126 |
||
127 |
lemma merge_vimage: |
|
128 |
assumes "I \<inter> J = {}" |
|
129 |
shows "(\<lambda>(x,y). merge I x J y) -` Pi\<^isub>E (I \<union> J) E = Pi I E \<times> Pi J E" |
|
130 |
using assms by (auto simp: restrict_Pi_cancel) |
|
131 |
||
132 |
lemma restrict_fupd[simp]: "i \<notin> I \<Longrightarrow> restrict (f (i := x)) I = restrict f I" |
|
133 |
by (auto simp: restrict_def intro!: ext) |
|
134 |
||
135 |
lemma merge_restrict[simp]: |
|
136 |
"merge I (restrict x I) J y = merge I x J y" |
|
137 |
"merge I x J (restrict y J) = merge I x J y" |
|
138 |
unfolding merge_def by (auto intro!: ext) |
|
139 |
||
140 |
lemma merge_x_x_eq_restrict[simp]: |
|
141 |
"merge I x J x = restrict x (I \<union> J)" |
|
142 |
unfolding merge_def by (auto intro!: ext) |
|
143 |
||
144 |
lemma Pi_fupd_iff: "i \<in> I \<Longrightarrow> f \<in> Pi I (B(i := A)) \<longleftrightarrow> f \<in> Pi (I - {i}) B \<and> f i \<in> A" |
|
145 |
apply auto |
|
146 |
apply (drule_tac x=x in Pi_mem) |
|
147 |
apply (simp_all split: split_if_asm) |
|
148 |
apply (drule_tac x=i in Pi_mem) |
|
149 |
apply (auto dest!: Pi_mem) |
|
150 |
done |
|
151 |
||
152 |
lemma Pi_UN: |
|
153 |
fixes A :: "nat \<Rightarrow> 'i \<Rightarrow> 'a set" |
|
154 |
assumes "finite I" and mono: "\<And>i n m. i \<in> I \<Longrightarrow> n \<le> m \<Longrightarrow> A n i \<subseteq> A m i" |
|
155 |
shows "(\<Union>n. Pi I (A n)) = (\<Pi> i\<in>I. \<Union>n. A n i)" |
|
156 |
proof (intro set_eqI iffI) |
|
157 |
fix f assume "f \<in> (\<Pi> i\<in>I. \<Union>n. A n i)" |
|
158 |
then have "\<forall>i\<in>I. \<exists>n. f i \<in> A n i" by auto |
|
159 |
from bchoice[OF this] obtain n where n: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> (A (n i) i)" by auto |
|
160 |
obtain k where k: "\<And>i. i \<in> I \<Longrightarrow> n i \<le> k" |
|
161 |
using `finite I` finite_nat_set_iff_bounded_le[of "n`I"] by auto |
|
162 |
have "f \<in> Pi I (A k)" |
|
163 |
proof (intro Pi_I) |
|
164 |
fix i assume "i \<in> I" |
|
165 |
from mono[OF this, of "n i" k] k[OF this] n[OF this] |
|
166 |
show "f i \<in> A k i" by auto |
|
167 |
qed |
|
168 |
then show "f \<in> (\<Union>n. Pi I (A n))" by auto |
|
169 |
qed auto |
|
170 |
||
171 |
lemma PiE_cong: |
|
172 |
assumes "\<And>i. i\<in>I \<Longrightarrow> A i = B i" |
|
173 |
shows "Pi\<^isub>E I A = Pi\<^isub>E I B" |
|
174 |
using assms by (auto intro!: Pi_cong) |
|
175 |
||
176 |
lemma restrict_upd[simp]: |
|
177 |
"i \<notin> I \<Longrightarrow> (restrict f I)(i := y) = restrict (f(i := y)) (insert i I)" |
|
178 |
by (auto simp: fun_eq_iff) |
|
179 |
||
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
|
180 |
lemma Pi_eq_subset: |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
181 |
assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
182 |
assumes eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and "i \<in> I" |
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
|
183 |
shows "F i \<subseteq> F' i" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
184 |
proof |
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
|
185 |
fix x assume "x \<in> F i" |
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
|
186 |
with ne have "\<forall>j. \<exists>y. ((j \<in> I \<longrightarrow> y \<in> F j \<and> (i = j \<longrightarrow> x = y)) \<and> (j \<notin> I \<longrightarrow> y = undefined))" by auto |
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
|
187 |
from choice[OF this] guess f .. note f = this |
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
|
188 |
then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def) |
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
|
189 |
then have "f \<in> Pi\<^isub>E I F'" using assms by simp |
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
|
190 |
then show "x \<in> F' i" using f `i \<in> I` by auto |
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
|
191 |
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
|
192 |
|
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
|
193 |
lemma Pi_eq_iff_not_empty: |
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
|
194 |
assumes ne: "\<And>i. i \<in> I \<Longrightarrow> F i \<noteq> {}" "\<And>i. i \<in> I \<Longrightarrow> F' i \<noteq> {}" |
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
|
195 |
shows "Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i)" |
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
|
196 |
proof (intro iffI ballI) |
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
|
197 |
fix i assume eq: "Pi\<^isub>E I F = Pi\<^isub>E I F'" and i: "i \<in> I" |
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
|
198 |
show "F i = F' i" |
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
|
199 |
using Pi_eq_subset[of I F F', OF ne eq i] |
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
|
200 |
using Pi_eq_subset[of I F' F, OF ne(2,1) eq[symmetric] i] |
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
|
201 |
by auto |
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
|
202 |
qed auto |
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
|
203 |
|
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
|
204 |
lemma Pi_eq_empty_iff: |
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
|
205 |
"Pi\<^isub>E I F = {} \<longleftrightarrow> (\<exists>i\<in>I. F i = {})" |
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
|
206 |
proof |
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
|
207 |
assume "Pi\<^isub>E I F = {}" |
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
|
208 |
show "\<exists>i\<in>I. F i = {}" |
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
|
209 |
proof (rule ccontr) |
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
|
210 |
assume "\<not> ?thesis" |
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
|
211 |
then have "\<forall>i. \<exists>y. (i \<in> I \<longrightarrow> y \<in> F i) \<and> (i \<notin> I \<longrightarrow> y = undefined)" by auto |
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
|
212 |
from choice[OF this] guess f .. |
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
|
213 |
then have "f \<in> Pi\<^isub>E I F" by (auto simp: extensional_def) |
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
|
214 |
with `Pi\<^isub>E I F = {}` show False by auto |
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
|
215 |
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
|
216 |
qed auto |
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
|
217 |
|
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
|
218 |
lemma Pi_eq_iff: |
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
|
219 |
"Pi\<^isub>E I F = Pi\<^isub>E I F' \<longleftrightarrow> (\<forall>i\<in>I. F i = F' i) \<or> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))" |
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
|
220 |
proof (intro iffI disjCI) |
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
|
221 |
assume eq[simp]: "Pi\<^isub>E I F = Pi\<^isub>E I F'" |
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
|
222 |
assume "\<not> ((\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {}))" |
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
|
223 |
then have "(\<forall>i\<in>I. F i \<noteq> {}) \<and> (\<forall>i\<in>I. F' i \<noteq> {})" |
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
|
224 |
using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto |
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
|
225 |
with Pi_eq_iff_not_empty[of I F F'] show "\<forall>i\<in>I. F i = F' i" by auto |
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
|
226 |
next |
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
|
227 |
assume "(\<forall>i\<in>I. F i = F' i) \<or> (\<exists>i\<in>I. F i = {}) \<and> (\<exists>i\<in>I. F' i = {})" |
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
|
228 |
then show "Pi\<^isub>E I F = Pi\<^isub>E I F'" |
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
|
229 |
using Pi_eq_empty_iff[of I F] Pi_eq_empty_iff[of I F'] by auto |
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
|
230 |
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
|
231 |
|
40859 | 232 |
section "Finite product spaces" |
233 |
||
234 |
section "Products" |
|
235 |
||
236 |
locale product_sigma_algebra = |
|
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
|
237 |
fixes M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme" |
40859 | 238 |
assumes sigma_algebras: "\<And>i. sigma_algebra (M i)" |
239 |
||
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
|
240 |
locale finite_product_sigma_algebra = product_sigma_algebra M |
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
|
241 |
for M :: "'i \<Rightarrow> ('a, 'b) measure_space_scheme" + |
40859 | 242 |
fixes I :: "'i set" |
243 |
assumes finite_index: "finite I" |
|
244 |
||
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
|
245 |
definition |
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
|
246 |
"product_algebra_generator I M = \<lparr> space = (\<Pi>\<^isub>E i \<in> I. space (M i)), |
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
|
247 |
sets = Pi\<^isub>E I ` (\<Pi> i \<in> I. sets (M i)), |
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
|
248 |
measure = \<lambda>A. (\<Prod>i\<in>I. measure (M i) ((SOME F. A = Pi\<^isub>E I F) i)) \<rparr>" |
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
|
249 |
|
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
|
250 |
definition product_algebra_def: |
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
|
251 |
"Pi\<^isub>M I M = sigma (product_algebra_generator I M) |
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
|
252 |
\<lparr> measure := (SOME \<mu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<mu> \<rparr>) \<and> |
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
|
253 |
(\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. measure (M i) (A i))))\<rparr>" |
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
|
254 |
|
40859 | 255 |
syntax |
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
|
256 |
"_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => |
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
|
257 |
('i => 'a, 'b) measure_space_scheme" ("(3PIM _:_./ _)" 10) |
40859 | 258 |
|
259 |
syntax (xsymbols) |
|
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
|
260 |
"_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => |
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
|
261 |
('i => 'a, 'b) measure_space_scheme" ("(3\<Pi>\<^isub>M _\<in>_./ _)" 10) |
40859 | 262 |
|
263 |
syntax (HTML output) |
|
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
|
264 |
"_PiM" :: "[pttrn, 'i set, ('a, 'b) measure_space_scheme] => |
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
|
265 |
('i => 'a, 'b) measure_space_scheme" ("(3\<Pi>\<^isub>M _\<in>_./ _)" 10) |
40859 | 266 |
|
267 |
translations |
|
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
|
268 |
"PIM x:I. M" == "CONST Pi\<^isub>M I (%x. M)" |
40859 | 269 |
|
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
|
270 |
abbreviation (in finite_product_sigma_algebra) "G \<equiv> product_algebra_generator I M" |
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
|
271 |
abbreviation (in finite_product_sigma_algebra) "P \<equiv> Pi\<^isub>M I M" |
40859 | 272 |
|
273 |
sublocale product_sigma_algebra \<subseteq> M: sigma_algebra "M i" for i by (rule sigma_algebras) |
|
274 |
||
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
|
275 |
lemma sigma_into_space: |
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
|
276 |
assumes "sets M \<subseteq> Pow (space M)" |
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
|
277 |
shows "sets (sigma M) \<subseteq> Pow (space M)" |
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
|
278 |
using sigma_sets_into_sp[OF assms] unfolding sigma_def by auto |
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
|
279 |
|
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
|
280 |
lemma (in product_sigma_algebra) product_algebra_generator_into_space: |
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
|
281 |
"sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))" |
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
|
282 |
using M.sets_into_space unfolding product_algebra_generator_def |
40859 | 283 |
by auto blast |
284 |
||
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
|
285 |
lemma (in product_sigma_algebra) product_algebra_into_space: |
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
|
286 |
"sets (Pi\<^isub>M I M) \<subseteq> Pow (space (Pi\<^isub>M I M))" |
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
|
287 |
using product_algebra_generator_into_space |
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
|
288 |
by (auto intro!: sigma_into_space simp add: product_algebra_def) |
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
|
289 |
|
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
|
290 |
lemma (in product_sigma_algebra) sigma_algebra_product_algebra: "sigma_algebra (Pi\<^isub>M I M)" |
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
|
291 |
using product_algebra_generator_into_space unfolding product_algebra_def |
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
|
292 |
by (rule sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) simp_all |
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
|
293 |
|
40859 | 294 |
sublocale finite_product_sigma_algebra \<subseteq> sigma_algebra P |
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
|
295 |
using sigma_algebra_product_algebra . |
40859 | 296 |
|
41095 | 297 |
lemma product_algebraE: |
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
|
298 |
assumes "A \<in> sets (product_algebra_generator I M)" |
41095 | 299 |
obtains E where "A = Pi\<^isub>E I E" "E \<in> (\<Pi> i\<in>I. sets (M i))" |
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
|
300 |
using assms unfolding product_algebra_generator_def by auto |
41095 | 301 |
|
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
|
302 |
lemma product_algebra_generatorI[intro]: |
41095 | 303 |
assumes "E \<in> (\<Pi> i\<in>I. sets (M i))" |
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
|
304 |
shows "Pi\<^isub>E I E \<in> sets (product_algebra_generator I M)" |
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 |
using assms unfolding product_algebra_generator_def by auto |
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
|
306 |
|
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
|
307 |
lemma space_product_algebra_generator[simp]: |
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
|
308 |
"space (product_algebra_generator I M) = Pi\<^isub>E I (\<lambda>i. space (M i))" |
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
|
309 |
unfolding product_algebra_generator_def by simp |
41095 | 310 |
|
40859 | 311 |
lemma space_product_algebra[simp]: |
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
|
312 |
"space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E i\<in>I. space (M i))" |
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
|
313 |
unfolding product_algebra_def product_algebra_generator_def by simp |
40859 | 314 |
|
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
|
315 |
lemma sets_product_algebra: |
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
|
316 |
"sets (Pi\<^isub>M I M) = sets (sigma (product_algebra_generator I M))" |
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
|
317 |
unfolding product_algebra_def sigma_def by simp |
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
|
318 |
|
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
|
319 |
lemma product_algebra_generator_sets_into_space: |
41095 | 320 |
assumes "\<And>i. i\<in>I \<Longrightarrow> sets (M i) \<subseteq> Pow (space (M i))" |
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
|
321 |
shows "sets (product_algebra_generator I M) \<subseteq> Pow (space (product_algebra_generator I M))" |
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
|
322 |
using assms by (auto simp: product_algebra_generator_def) blast |
40859 | 323 |
|
324 |
lemma (in finite_product_sigma_algebra) in_P[simp, intro]: |
|
325 |
"\<lbrakk> \<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i) \<rbrakk> \<Longrightarrow> Pi\<^isub>E I A \<in> sets P" |
|
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
|
326 |
by (auto simp: sets_product_algebra) |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
327 |
|
42988 | 328 |
lemma Int_stable_product_algebra_generator: |
329 |
"(\<And>i. i \<in> I \<Longrightarrow> Int_stable (M i)) \<Longrightarrow> Int_stable (product_algebra_generator I M)" |
|
330 |
by (auto simp add: product_algebra_generator_def Int_stable_def PiE_Int Pi_iff) |
|
331 |
||
40859 | 332 |
section "Generating set generates also product algebra" |
333 |
||
334 |
lemma sigma_product_algebra_sigma_eq: |
|
335 |
assumes "finite I" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
336 |
assumes mono: "\<And>i. i \<in> I \<Longrightarrow> incseq (S i)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
337 |
assumes union: "\<And>i. i \<in> I \<Longrightarrow> (\<Union>j. S i j) = space (E i)" |
40859 | 338 |
assumes sets_into: "\<And>i. i \<in> I \<Longrightarrow> range (S i) \<subseteq> sets (E i)" |
339 |
and E: "\<And>i. sets (E i) \<subseteq> Pow (space (E i))" |
|
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
|
340 |
shows "sets (\<Pi>\<^isub>M i\<in>I. sigma (E i)) = sets (\<Pi>\<^isub>M i\<in>I. E i)" |
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
|
341 |
(is "sets ?S = sets ?E") |
40859 | 342 |
proof cases |
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 |
assume "I = {}" then show ?thesis |
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
|
344 |
by (simp add: product_algebra_def product_algebra_generator_def) |
40859 | 345 |
next |
346 |
assume "I \<noteq> {}" |
|
347 |
interpret E: sigma_algebra "sigma (E i)" for i |
|
348 |
using E by (rule sigma_algebra_sigma) |
|
349 |
have into_space[intro]: "\<And>i x A. A \<in> sets (E i) \<Longrightarrow> x i \<in> A \<Longrightarrow> x i \<in> space (E i)" |
|
350 |
using E by auto |
|
351 |
interpret G: sigma_algebra ?E |
|
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
|
352 |
unfolding product_algebra_def product_algebra_generator_def using E |
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
|
353 |
by (intro sigma_algebra.sigma_algebra_cong[OF sigma_algebra_sigma]) (auto dest: Pi_mem) |
40859 | 354 |
{ fix A i assume "i \<in> I" and A: "A \<in> sets (E i)" |
355 |
then have "(\<lambda>x. x i) -` A \<inter> space ?E = (\<Pi>\<^isub>E j\<in>I. if j = i then A else \<Union>n. S j n) \<inter> space ?E" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
356 |
using mono union unfolding incseq_Suc_iff space_product_algebra |
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
|
357 |
by (auto dest: Pi_mem) |
40859 | 358 |
also have "\<dots> = (\<Union>n. (\<Pi>\<^isub>E j\<in>I. if j = i then A else S j n))" |
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
|
359 |
unfolding space_product_algebra |
40859 | 360 |
apply simp |
361 |
apply (subst Pi_UN[OF `finite I`]) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
362 |
using mono[THEN incseqD] apply simp |
40859 | 363 |
apply (simp add: PiE_Int) |
364 |
apply (intro PiE_cong) |
|
365 |
using A sets_into by (auto intro!: 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
|
366 |
also have "\<dots> \<in> sets ?E" |
40859 | 367 |
using sets_into `A \<in> sets (E i)` |
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
|
368 |
unfolding sets_product_algebra sets_sigma |
40859 | 369 |
by (intro sigma_sets.Union) |
370 |
(auto simp: image_subset_iff intro!: sigma_sets.Basic) |
|
371 |
finally have "(\<lambda>x. x i) -` A \<inter> space ?E \<in> sets ?E" . } |
|
372 |
then have proj: |
|
373 |
"\<And>i. i\<in>I \<Longrightarrow> (\<lambda>x. x i) \<in> measurable ?E (sigma (E i))" |
|
374 |
using E by (subst G.measurable_iff_sigma) |
|
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
|
375 |
(auto simp: sets_product_algebra sets_sigma) |
40859 | 376 |
{ fix A assume A: "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (sigma (E i))" |
377 |
with proj have basic: "\<And>i. i \<in> I \<Longrightarrow> (\<lambda>x. x i) -` (A i) \<inter> space ?E \<in> sets ?E" |
|
378 |
unfolding measurable_def by simp |
|
379 |
have "Pi\<^isub>E I A = (\<Inter>i\<in>I. (\<lambda>x. x i) -` (A i) \<inter> space ?E)" |
|
380 |
using A E.sets_into_space `I \<noteq> {}` unfolding product_algebra_def by auto blast |
|
381 |
then have "Pi\<^isub>E I A \<in> sets ?E" |
|
382 |
using G.finite_INT[OF `finite I` `I \<noteq> {}` basic, of "\<lambda>i. i"] by simp } |
|
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
|
383 |
then have "sigma_sets (space ?E) (sets (product_algebra_generator I (\<lambda>i. sigma (E i)))) \<subseteq> sets ?E" |
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
|
384 |
by (intro G.sigma_sets_subset) (auto simp add: product_algebra_generator_def) |
40859 | 385 |
then have subset: "sets ?S \<subseteq> sets ?E" |
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
|
386 |
by (simp add: sets_sigma sets_product_algebra) |
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
|
387 |
show "sets ?S = sets ?E" |
40859 | 388 |
proof (intro set_eqI iffI) |
389 |
fix A assume "A \<in> sets ?E" then show "A \<in> sets ?S" |
|
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
|
390 |
unfolding sets_sigma sets_product_algebra |
40859 | 391 |
proof induct |
392 |
case (Basic A) then show ?case |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
393 |
by (auto simp: sets_sigma product_algebra_generator_def intro: sigma_sets.Basic) |
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
|
394 |
qed (auto intro: sigma_sets.intros simp: product_algebra_generator_def) |
40859 | 395 |
next |
396 |
fix A assume "A \<in> sets ?S" then show "A \<in> sets ?E" using subset by auto |
|
397 |
qed |
|
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
|
398 |
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
|
399 |
|
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
400 |
lemma product_algebraI[intro]: |
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
|
401 |
"E \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> Pi\<^isub>E I E \<in> sets (Pi\<^isub>M I M)" |
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 |
using assms unfolding product_algebra_def by (auto intro: product_algebra_generatorI) |
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
|
403 |
|
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
|
404 |
lemma (in product_sigma_algebra) measurable_component_update: |
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
|
405 |
assumes "x \<in> space (Pi\<^isub>M I M)" and "i \<notin> I" |
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
|
406 |
shows "(\<lambda>v. x(i := v)) \<in> measurable (M i) (Pi\<^isub>M (insert i I) M)" (is "?f \<in> _") |
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
|
407 |
unfolding product_algebra_def apply simp |
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
|
408 |
proof (intro measurable_sigma) |
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
|
409 |
let ?G = "product_algebra_generator (insert i I) M" |
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
|
410 |
show "sets ?G \<subseteq> Pow (space ?G)" using product_algebra_generator_into_space . |
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 |
show "?f \<in> space (M i) \<rightarrow> space ?G" |
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
|
412 |
using M.sets_into_space assms by auto |
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
|
413 |
fix A assume "A \<in> sets ?G" |
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
|
414 |
from product_algebraE[OF this] guess E . note E = this |
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
|
415 |
then have "?f -` A \<inter> space (M i) = E i \<or> ?f -` A \<inter> space (M i) = {}" |
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
|
416 |
using M.sets_into_space assms by auto |
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
|
417 |
then show "?f -` A \<inter> space (M i) \<in> sets (M i)" |
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
|
418 |
using E by (auto intro!: product_algebraI) |
40859 | 419 |
qed |
420 |
||
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
421 |
lemma (in product_sigma_algebra) measurable_add_dim: |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
422 |
assumes "i \<notin> I" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
423 |
shows "(\<lambda>(f, y). f(i := y)) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) (Pi\<^isub>M (insert i I) M)" |
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
|
424 |
proof - |
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
|
425 |
let ?f = "(\<lambda>(f, y). f(i := y))" and ?G = "product_algebra_generator (insert i I) M" |
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
|
426 |
interpret Ii: pair_sigma_algebra "Pi\<^isub>M I M" "M i" |
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
|
427 |
unfolding pair_sigma_algebra_def |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
428 |
by (intro sigma_algebra_product_algebra sigma_algebras conjI) |
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
|
429 |
have "?f \<in> measurable Ii.P (sigma ?G)" |
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
|
430 |
proof (rule Ii.measurable_sigma) |
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
|
431 |
show "sets ?G \<subseteq> Pow (space ?G)" |
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 |
using product_algebra_generator_into_space . |
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 |
show "?f \<in> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<rightarrow> space ?G" |
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 |
by (auto simp: space_pair_measure) |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
435 |
next |
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
|
436 |
fix A assume "A \<in> sets ?G" |
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
|
437 |
then obtain F where "A = Pi\<^isub>E (insert i I) F" |
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
|
438 |
and F: "\<And>j. j \<in> I \<Longrightarrow> F j \<in> sets (M j)" "F i \<in> sets (M i)" |
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 |
by (auto elim!: product_algebraE) |
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
|
440 |
then have "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) = Pi\<^isub>E I F \<times> (F i)" |
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
|
441 |
using sets_into_space `i \<notin> I` |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
442 |
by (auto simp add: space_pair_measure) blast+ |
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
|
443 |
then show "?f -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M M i) \<in> sets (Pi\<^isub>M I M \<Otimes>\<^isub>M M i)" |
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
|
444 |
using F by (auto intro!: pair_measureI) |
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
|
445 |
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
|
446 |
then show ?thesis |
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
|
447 |
by (simp add: product_algebra_def) |
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
|
448 |
qed |
41095 | 449 |
|
450 |
lemma (in product_sigma_algebra) measurable_merge: |
|
451 |
assumes [simp]: "I \<inter> J = {}" |
|
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
|
452 |
shows "(\<lambda>(x, y). merge I x J y) \<in> measurable (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)" |
40859 | 453 |
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
|
454 |
let ?I = "Pi\<^isub>M I M" and ?J = "Pi\<^isub>M J M" |
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
|
455 |
interpret P: sigma_algebra "?I \<Otimes>\<^isub>M ?J" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
456 |
by (intro sigma_algebra_pair_measure product_algebra_into_space) |
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
|
457 |
let ?f = "\<lambda>(x, y). merge I x J y" |
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
|
458 |
let ?G = "product_algebra_generator (I \<union> J) M" |
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
|
459 |
have "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (sigma ?G)" |
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
|
460 |
proof (rule P.measurable_sigma) |
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
|
461 |
fix A assume "A \<in> sets ?G" |
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
|
462 |
from product_algebraE[OF this] |
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
|
463 |
obtain E where E: "A = Pi\<^isub>E (I \<union> J) E" "E \<in> (\<Pi> i\<in>I \<union> J. sets (M i))" . |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
464 |
then have *: "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) = Pi\<^isub>E I E \<times> Pi\<^isub>E J E" |
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
|
465 |
using sets_into_space `I \<inter> J = {}` |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
466 |
by (auto simp add: space_pair_measure) fast+ |
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
|
467 |
then show "?f -` A \<inter> space (?I \<Otimes>\<^isub>M ?J) \<in> sets (?I \<Otimes>\<^isub>M ?J)" |
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
|
468 |
using E unfolding * by (auto intro!: pair_measureI in_sigma product_algebra_generatorI |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
469 |
simp: product_algebra_def) |
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
|
470 |
qed (insert product_algebra_generator_into_space, auto simp: space_pair_measure) |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
471 |
then show "?f \<in> measurable (?I \<Otimes>\<^isub>M ?J) (Pi\<^isub>M (I \<union> J) M)" |
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
|
472 |
unfolding product_algebra_def[of "I \<union> J"] by simp |
40859 | 473 |
qed |
474 |
||
41095 | 475 |
lemma (in product_sigma_algebra) measurable_component_singleton: |
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
|
476 |
assumes "i \<in> I" shows "(\<lambda>x. x i) \<in> measurable (Pi\<^isub>M I M) (M i)" |
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
|
477 |
proof (unfold measurable_def, intro CollectI conjI ballI) |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
478 |
fix A assume "A \<in> sets (M i)" |
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
|
479 |
then have "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) = (\<Pi>\<^isub>E j\<in>I. if i = j then A else space (M j))" |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
43920
diff
changeset
|
480 |
using M.sets_into_space `i \<in> I` by (fastforce dest: Pi_mem split: split_if_asm) |
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
|
481 |
then show "(\<lambda>x. x i) -` A \<inter> space (Pi\<^isub>M I M) \<in> sets (Pi\<^isub>M I M)" |
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
|
482 |
using `A \<in> sets (M i)` by (auto intro!: product_algebraI) |
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
|
483 |
qed (insert `i \<in> I`, auto) |
41661 | 484 |
|
42988 | 485 |
lemma (in sigma_algebra) measurable_restrict: |
486 |
assumes I: "finite I" |
|
487 |
assumes "\<And>i. i \<in> I \<Longrightarrow> sets (N i) \<subseteq> Pow (space (N i))" |
|
488 |
assumes X: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> measurable M (N i)" |
|
489 |
shows "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> measurable M (Pi\<^isub>M I N)" |
|
490 |
unfolding product_algebra_def |
|
491 |
proof (simp, rule measurable_sigma) |
|
492 |
show "sets (product_algebra_generator I N) \<subseteq> Pow (space (product_algebra_generator I N))" |
|
493 |
by (rule product_algebra_generator_sets_into_space) fact |
|
494 |
show "(\<lambda>x. \<lambda>i\<in>I. X i x) \<in> space M \<rightarrow> space (product_algebra_generator I N)" |
|
495 |
using X by (auto simp: measurable_def) |
|
496 |
fix E assume "E \<in> sets (product_algebra_generator I N)" |
|
497 |
then obtain F where "E = Pi\<^isub>E I F" and F: "\<And>i. i \<in> I \<Longrightarrow> F i \<in> sets (N i)" |
|
498 |
by (auto simp: product_algebra_generator_def) |
|
499 |
then have "(\<lambda>x. \<lambda>i\<in>I. X i x) -` E \<inter> space M = (\<Inter>i\<in>I. X i -` F i \<inter> space M) \<inter> space M" |
|
500 |
by (auto simp: Pi_iff) |
|
501 |
also have "\<dots> \<in> sets M" |
|
502 |
proof cases |
|
503 |
assume "I = {}" then show ?thesis by simp |
|
504 |
next |
|
505 |
assume "I \<noteq> {}" with X F I show ?thesis |
|
506 |
by (intro finite_INT measurable_sets Int) auto |
|
507 |
qed |
|
508 |
finally show "(\<lambda>x. \<lambda>i\<in>I. X i x) -` E \<inter> space M \<in> sets M" . |
|
509 |
qed |
|
510 |
||
40859 | 511 |
locale product_sigma_finite = |
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
|
512 |
fixes M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" |
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
|
513 |
assumes sigma_finite_measures: "\<And>i. sigma_finite_measure (M i)" |
40859 | 514 |
|
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
|
515 |
locale finite_product_sigma_finite = product_sigma_finite M |
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
|
516 |
for M :: "'i \<Rightarrow> ('a,'b) measure_space_scheme" + |
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
|
517 |
fixes I :: "'i set" assumes finite_index'[intro]: "finite I" |
40859 | 518 |
|
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
|
519 |
sublocale product_sigma_finite \<subseteq> M: sigma_finite_measure "M i" for i |
40859 | 520 |
by (rule sigma_finite_measures) |
521 |
||
522 |
sublocale product_sigma_finite \<subseteq> product_sigma_algebra |
|
523 |
by default |
|
524 |
||
525 |
sublocale finite_product_sigma_finite \<subseteq> finite_product_sigma_algebra |
|
526 |
by default (fact finite_index') |
|
527 |
||
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
|
528 |
lemma (in finite_product_sigma_finite) product_algebra_generator_measure: |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
529 |
assumes "Pi\<^isub>E I F \<in> sets G" |
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
|
530 |
shows "measure G (Pi\<^isub>E I F) = (\<Prod>i\<in>I. M.\<mu> i (F i))" |
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
|
531 |
proof cases |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
532 |
assume ne: "\<forall>i\<in>I. F i \<noteq> {}" |
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
|
533 |
have "\<forall>i\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') i = F i" |
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
|
534 |
by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) |
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
|
535 |
(insert ne, auto simp: Pi_eq_iff) |
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
|
536 |
then show ?thesis |
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
|
537 |
unfolding product_algebra_generator_def by simp |
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
|
538 |
next |
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
|
539 |
assume empty: "\<not> (\<forall>j\<in>I. F j \<noteq> {})" |
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
|
540 |
then have "(\<Prod>j\<in>I. M.\<mu> j (F j)) = 0" |
43920 | 541 |
by (auto simp: setprod_ereal_0 intro!: bexI) |
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
|
542 |
moreover |
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
|
543 |
have "\<exists>j\<in>I. (SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j = {}" |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
544 |
by (rule someI2[where P="\<lambda>F'. Pi\<^isub>E I F = Pi\<^isub>E I F'"]) |
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
|
545 |
(insert empty, auto simp: Pi_eq_empty_iff[symmetric]) |
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
|
546 |
then have "(\<Prod>j\<in>I. M.\<mu> j ((SOME F'. Pi\<^isub>E I F = Pi\<^isub>E I F') j)) = 0" |
43920 | 547 |
by (auto simp: setprod_ereal_0 intro!: bexI) |
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
|
548 |
ultimately show ?thesis |
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
|
549 |
unfolding product_algebra_generator_def by simp |
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
|
550 |
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
|
551 |
|
40859 | 552 |
lemma (in finite_product_sigma_finite) sigma_finite_pairs: |
553 |
"\<exists>F::'i \<Rightarrow> nat \<Rightarrow> 'a set. |
|
554 |
(\<forall>i\<in>I. range (F i) \<subseteq> sets (M i)) \<and> |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
555 |
(\<forall>k. \<forall>i\<in>I. \<mu> i (F i k) \<noteq> \<infinity>) \<and> incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k) \<and> |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
556 |
(\<Union>k. \<Pi>\<^isub>E i\<in>I. F i k) = space G" |
40859 | 557 |
proof - |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
558 |
have "\<forall>i::'i. \<exists>F::nat \<Rightarrow> 'a set. range F \<subseteq> sets (M i) \<and> incseq F \<and> (\<Union>i. F i) = space (M i) \<and> (\<forall>k. \<mu> i (F k) \<noteq> \<infinity>)" |
40859 | 559 |
using M.sigma_finite_up by simp |
560 |
from choice[OF this] guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" .. |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
561 |
then have F: "\<And>i. range (F i) \<subseteq> sets (M i)" "\<And>i. incseq (F i)" "\<And>i. (\<Union>j. F i j) = space (M i)" "\<And>i k. \<mu> i (F i k) \<noteq> \<infinity>" |
40859 | 562 |
by auto |
563 |
let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I. F i k" |
|
564 |
note space_product_algebra[simp] |
|
565 |
show ?thesis |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
566 |
proof (intro exI[of _ F] conjI allI incseq_SucI set_eqI iffI ballI) |
40859 | 567 |
fix i show "range (F i) \<subseteq> sets (M i)" by fact |
568 |
next |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
569 |
fix i k show "\<mu> i (F i k) \<noteq> \<infinity>" by fact |
40859 | 570 |
next |
571 |
fix A assume "A \<in> (\<Union>i. ?F i)" then show "A \<in> space G" |
|
41831 | 572 |
using `\<And>i. range (F i) \<subseteq> sets (M i)` M.sets_into_space |
573 |
by (force simp: image_subset_iff) |
|
40859 | 574 |
next |
575 |
fix f assume "f \<in> space G" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
576 |
with Pi_UN[OF finite_index, of "\<lambda>k i. F i k"] F |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
577 |
show "f \<in> (\<Union>i. ?F i)" by (auto simp: incseq_def) |
40859 | 578 |
next |
579 |
fix i show "?F i \<subseteq> ?F (Suc i)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
580 |
using `\<And>i. incseq (F i)`[THEN incseq_SucD] by auto |
40859 | 581 |
qed |
582 |
qed |
|
583 |
||
41831 | 584 |
lemma sets_pair_cancel_measure[simp]: |
585 |
"sets (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) = sets (M1 \<Otimes>\<^isub>M M2)" |
|
586 |
"sets (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) = sets (M1 \<Otimes>\<^isub>M M2)" |
|
587 |
unfolding pair_measure_def pair_measure_generator_def sets_sigma |
|
588 |
by simp_all |
|
589 |
||
590 |
lemma measurable_pair_cancel_measure[simp]: |
|
591 |
"measurable (M1\<lparr>measure := m1\<rparr> \<Otimes>\<^isub>M M2) M = measurable (M1 \<Otimes>\<^isub>M M2) M" |
|
592 |
"measurable (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m2\<rparr>) M = measurable (M1 \<Otimes>\<^isub>M M2) M" |
|
593 |
"measurable M (M1\<lparr>measure := m3\<rparr> \<Otimes>\<^isub>M M2) = measurable M (M1 \<Otimes>\<^isub>M M2)" |
|
594 |
"measurable M (M1 \<Otimes>\<^isub>M M2\<lparr>measure := m4\<rparr>) = measurable M (M1 \<Otimes>\<^isub>M M2)" |
|
595 |
unfolding measurable_def by (auto simp add: space_pair_measure) |
|
596 |
||
40859 | 597 |
lemma (in product_sigma_finite) product_measure_exists: |
598 |
assumes "finite I" |
|
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
|
599 |
shows "\<exists>\<nu>. sigma_finite_measure (sigma (product_algebra_generator I M) \<lparr> measure := \<nu> \<rparr>) \<and> |
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
|
600 |
(\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i)))" |
40859 | 601 |
using `finite I` proof induct |
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
|
602 |
case empty |
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
|
603 |
interpret finite_product_sigma_finite M "{}" by default simp |
43920 | 604 |
let ?\<nu> = "(\<lambda>A. if A = {} then 0 else 1) :: 'd set \<Rightarrow> ereal" |
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
|
605 |
show ?case |
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
|
606 |
proof (intro exI conjI ballI) |
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
|
607 |
have "sigma_algebra (sigma G \<lparr>measure := ?\<nu>\<rparr>)" |
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
|
608 |
by (rule sigma_algebra_cong) (simp_all add: product_algebra_def) |
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
|
609 |
then have "measure_space (sigma G\<lparr>measure := ?\<nu>\<rparr>)" |
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
|
610 |
by (rule finite_additivity_sufficient) |
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
|
611 |
(simp_all add: positive_def additive_def sets_sigma |
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
|
612 |
product_algebra_generator_def image_constant) |
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
|
613 |
then show "sigma_finite_measure (sigma G\<lparr>measure := ?\<nu>\<rparr>)" |
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
|
614 |
by (auto intro!: exI[of _ "\<lambda>i. {\<lambda>_. undefined}"] |
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
|
615 |
simp: sigma_finite_measure_def sigma_finite_measure_axioms_def |
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
|
616 |
product_algebra_generator_def) |
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
|
617 |
qed auto |
40859 | 618 |
next |
619 |
case (insert i I) |
|
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
|
620 |
interpret finite_product_sigma_finite M I by default fact |
40859 | 621 |
have "finite (insert i I)" using `finite I` by auto |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
622 |
interpret I': finite_product_sigma_finite M "insert i I" by default fact |
40859 | 623 |
from insert obtain \<nu> where |
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
|
624 |
prod: "\<And>A. A \<in> (\<Pi> i\<in>I. sets (M i)) \<Longrightarrow> \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" and |
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
|
625 |
"sigma_finite_measure (sigma G\<lparr> measure := \<nu> \<rparr>)" by auto |
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
|
626 |
then interpret I: sigma_finite_measure "P\<lparr> measure := \<nu>\<rparr>" unfolding product_algebra_def by simp |
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
|
627 |
interpret P: pair_sigma_finite "P\<lparr> measure := \<nu>\<rparr>" "M i" .. |
41661 | 628 |
let ?h = "(\<lambda>(f, y). f(i := y))" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
629 |
let ?\<nu> = "\<lambda>A. P.\<mu> (?h -` A \<inter> space P.P)" |
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
|
630 |
have I': "sigma_algebra (I'.P\<lparr> measure := ?\<nu> \<rparr>)" |
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
|
631 |
by (rule I'.sigma_algebra_cong) simp_all |
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
|
632 |
interpret I'': measure_space "I'.P\<lparr> measure := ?\<nu> \<rparr>" |
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
|
633 |
using measurable_add_dim[OF `i \<notin> I`] |
41831 | 634 |
by (intro P.measure_space_vimage[OF I']) (auto simp add: measure_preserving_def) |
40859 | 635 |
show ?case |
636 |
proof (intro exI[of _ ?\<nu>] conjI ballI) |
|
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
|
637 |
let "?m A" = "measure (Pi\<^isub>M I M\<lparr>measure := \<nu>\<rparr> \<Otimes>\<^isub>M M i) (?h -` A \<inter> space P.P)" |
40859 | 638 |
{ fix A assume A: "A \<in> (\<Pi> i\<in>insert i I. sets (M i))" |
41661 | 639 |
then have *: "?h -` Pi\<^isub>E (insert i I) A \<inter> space P.P = Pi\<^isub>E I A \<times> A i" |
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
|
640 |
using `i \<notin> I` M.sets_into_space by (auto simp: space_pair_measure space_product_algebra) blast |
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
|
641 |
show "?m (Pi\<^isub>E (insert i I) A) = (\<Prod>i\<in>insert i I. M.\<mu> i (A i))" |
41661 | 642 |
unfolding * using A |
40859 | 643 |
apply (subst P.pair_measure_times) |
44890
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
43920
diff
changeset
|
644 |
using A apply fastforce |
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
nipkow
parents:
43920
diff
changeset
|
645 |
using A apply fastforce |
41661 | 646 |
using `i \<notin> I` `finite I` prod[of A] A by (auto simp: ac_simps) } |
40859 | 647 |
note product = this |
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
|
648 |
have *: "sigma I'.G\<lparr> measure := ?\<nu> \<rparr> = I'.P\<lparr> measure := ?\<nu> \<rparr>" |
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
|
649 |
by (simp add: product_algebra_def) |
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
|
650 |
show "sigma_finite_measure (sigma I'.G\<lparr> measure := ?\<nu> \<rparr>)" |
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
|
651 |
proof (unfold *, default, simp) |
40859 | 652 |
from I'.sigma_finite_pairs guess F :: "'i \<Rightarrow> nat \<Rightarrow> 'a set" .. |
653 |
then have F: "\<And>j. j \<in> insert i I \<Longrightarrow> range (F j) \<subseteq> sets (M j)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
654 |
"incseq (\<lambda>k. \<Pi>\<^isub>E j \<in> insert i I. F j k)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
655 |
"(\<Union>k. \<Pi>\<^isub>E j \<in> insert i I. F j k) = space I'.G" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
656 |
"\<And>k. \<And>j. j \<in> insert i I \<Longrightarrow> \<mu> j (F j k) \<noteq> \<infinity>" |
40859 | 657 |
by blast+ |
658 |
let "?F k" = "\<Pi>\<^isub>E j \<in> insert i I. F j k" |
|
659 |
show "\<exists>F::nat \<Rightarrow> ('i \<Rightarrow> 'a) set. range F \<subseteq> sets I'.P \<and> |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
660 |
(\<Union>i. F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i)) \<and> (\<forall>i. ?m (F i) \<noteq> \<infinity>)" |
40859 | 661 |
proof (intro exI[of _ ?F] conjI allI) |
662 |
show "range ?F \<subseteq> sets I'.P" using F(1) by auto |
|
663 |
next |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
664 |
from F(3) show "(\<Union>i. ?F i) = (\<Pi>\<^isub>E i\<in>insert i I. space (M i))" by simp |
40859 | 665 |
next |
666 |
fix j |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
667 |
have "\<And>k. k \<in> insert i I \<Longrightarrow> 0 \<le> measure (M k) (F k j)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
668 |
using F(1) by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
669 |
with F `finite I` setprod_PInf[of "insert i I", OF this] show "?\<nu> (?F j) \<noteq> \<infinity>" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
670 |
by (subst product) auto |
40859 | 671 |
qed |
672 |
qed |
|
673 |
qed |
|
674 |
qed |
|
675 |
||
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
|
676 |
sublocale finite_product_sigma_finite \<subseteq> sigma_finite_measure P |
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
|
677 |
unfolding product_algebra_def |
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
|
678 |
using product_measure_exists[OF finite_index] |
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
|
679 |
by (rule someI2_ex) auto |
40859 | 680 |
|
681 |
lemma (in finite_product_sigma_finite) measure_times: |
|
682 |
assumes "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets (M i)" |
|
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
|
683 |
shows "\<mu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" |
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
|
684 |
unfolding product_algebra_def |
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
|
685 |
using product_measure_exists[OF finite_index] |
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
|
686 |
proof (rule someI2_ex, elim conjE) |
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
|
687 |
fix \<nu> assume *: "\<forall>A\<in>\<Pi> i\<in>I. sets (M i). \<nu> (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" |
40859 | 688 |
have "Pi\<^isub>E I A = Pi\<^isub>E I (\<lambda>i\<in>I. A i)" by (auto dest: Pi_mem) |
689 |
then have "\<nu> (Pi\<^isub>E I A) = \<nu> (Pi\<^isub>E I (\<lambda>i\<in>I. A i))" by simp |
|
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
|
690 |
also have "\<dots> = (\<Prod>i\<in>I. M.\<mu> i ((\<lambda>i\<in>I. A i) i))" using assms * by auto |
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
|
691 |
finally show "measure (sigma G\<lparr>measure := \<nu>\<rparr>) (Pi\<^isub>E I A) = (\<Prod>i\<in>I. M.\<mu> i (A i))" |
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
|
692 |
by (simp add: sigma_def) |
40859 | 693 |
qed |
41096 | 694 |
|
695 |
lemma (in product_sigma_finite) product_measure_empty[simp]: |
|
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
|
696 |
"measure (Pi\<^isub>M {} M) {\<lambda>x. undefined} = 1" |
41096 | 697 |
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
|
698 |
interpret finite_product_sigma_finite M "{}" by default auto |
41096 | 699 |
from measure_times[of "\<lambda>x. {}"] show ?thesis by simp |
700 |
qed |
|
701 |
||
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
|
702 |
lemma (in finite_product_sigma_algebra) P_empty: |
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
|
703 |
assumes "I = {}" |
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
|
704 |
shows "space P = {\<lambda>k. undefined}" "sets P = { {}, {\<lambda>k. undefined} }" |
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
|
705 |
unfolding product_algebra_def product_algebra_generator_def `I = {}` |
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
|
706 |
by (simp_all add: sigma_def image_constant) |
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
|
707 |
|
40859 | 708 |
lemma (in product_sigma_finite) positive_integral_empty: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
709 |
assumes pos: "0 \<le> f (\<lambda>k. undefined)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
710 |
shows "integral\<^isup>P (Pi\<^isub>M {} M) f = f (\<lambda>k. undefined)" |
40859 | 711 |
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
|
712 |
interpret finite_product_sigma_finite M "{}" by default (fact finite.emptyI) |
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
|
713 |
have "\<And>A. measure (Pi\<^isub>M {} M) (Pi\<^isub>E {} A) = 1" |
40859 | 714 |
using assms by (subst measure_times) auto |
715 |
then show ?thesis |
|
40873 | 716 |
unfolding positive_integral_def simple_function_def simple_integral_def_raw |
40859 | 717 |
proof (simp add: P_empty, intro antisym) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
718 |
show "f (\<lambda>k. undefined) \<le> (SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined))" |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44890
diff
changeset
|
719 |
by (intro SUP_upper) (auto simp: le_fun_def split: split_max) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
720 |
show "(SUP f:{g. g \<le> max 0 \<circ> f}. f (\<lambda>k. undefined)) \<le> f (\<lambda>k. undefined)" using pos |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44890
diff
changeset
|
721 |
by (intro SUP_least) (auto simp: le_fun_def simp: max_def split: split_if_asm) |
40859 | 722 |
qed |
723 |
qed |
|
724 |
||
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
725 |
lemma (in product_sigma_finite) measure_fold: |
40859 | 726 |
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J" |
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
|
727 |
assumes A: "A \<in> sets (Pi\<^isub>M (I \<union> J) M)" |
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
728 |
shows "measure (Pi\<^isub>M (I \<union> J) M) A = |
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
729 |
measure (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) ((\<lambda>(x,y). merge I x J y) -` A \<inter> space (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M))" |
40859 | 730 |
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
|
731 |
interpret I: finite_product_sigma_finite M I by default fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
732 |
interpret J: finite_product_sigma_finite M J by default fact |
40859 | 733 |
have "finite (I \<union> J)" using fin by auto |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
734 |
interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
735 |
interpret P: pair_sigma_finite I.P J.P by default |
41661 | 736 |
let ?g = "\<lambda>(x,y). merge I x J y" |
737 |
let "?X S" = "?g -` S \<inter> space P.P" |
|
738 |
from IJ.sigma_finite_pairs obtain F where |
|
739 |
F: "\<And>i. i\<in> I \<union> J \<Longrightarrow> range (F i) \<subseteq> sets (M i)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
740 |
"incseq (\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
741 |
"(\<Union>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k) = space IJ.G" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
742 |
"\<And>k. \<forall>i\<in>I\<union>J. \<mu> i (F i k) \<noteq> \<infinity>" |
41661 | 743 |
by auto |
744 |
let ?F = "\<lambda>k. \<Pi>\<^isub>E i\<in>I \<union> J. F i k" |
|
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
745 |
show "IJ.\<mu> A = P.\<mu> (?X A)" |
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
746 |
proof (rule measure_unique_Int_stable_vimage) |
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
747 |
show "measure_space IJ.P" "measure_space P.P" by default |
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
748 |
show "sets (sigma IJ.G) = sets IJ.P" "space IJ.G = space IJ.P" "A \<in> sets (sigma IJ.G)" |
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
|
749 |
using A unfolding product_algebra_def by auto |
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
750 |
next |
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
|
751 |
show "Int_stable IJ.G" |
42988 | 752 |
by (rule Int_stable_product_algebra_generator) |
753 |
(auto simp: Int_stable_def) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
754 |
show "range ?F \<subseteq> sets IJ.G" using F |
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
|
755 |
by (simp add: image_subset_iff product_algebra_def |
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
|
756 |
product_algebra_generator_def) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
757 |
show "incseq ?F" "(\<Union>i. ?F i) = space IJ.G " by fact+ |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
758 |
next |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
759 |
fix k |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
760 |
have "\<And>j. j \<in> I \<union> J \<Longrightarrow> 0 \<le> measure (M j) (F j k)" |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
761 |
using F(1) by auto |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
762 |
with F `finite I` setprod_PInf[of "I \<union> J", OF this] |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
763 |
show "IJ.\<mu> (?F k) \<noteq> \<infinity>" by (subst IJ.measure_times) auto |
41661 | 764 |
next |
765 |
fix A assume "A \<in> sets IJ.G" |
|
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
766 |
then obtain F where A: "A = Pi\<^isub>E (I \<union> J) F" |
41661 | 767 |
and F: "\<And>i. i \<in> I \<union> J \<Longrightarrow> F i \<in> sets (M i)" |
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
|
768 |
by (auto simp: product_algebra_generator_def) |
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
769 |
then have X: "?X A = (Pi\<^isub>E I F \<times> Pi\<^isub>E J 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
|
770 |
using sets_into_space by (auto simp: space_pair_measure) blast+ |
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
|
771 |
then have "P.\<mu> (?X A) = (\<Prod>i\<in>I. \<mu> i (F i)) * (\<Prod>i\<in>J. \<mu> i (F i))" |
41661 | 772 |
using `finite J` `finite I` F |
773 |
by (simp add: P.pair_measure_times I.measure_times J.measure_times) |
|
774 |
also have "\<dots> = (\<Prod>i\<in>I \<union> J. \<mu> i (F i))" |
|
775 |
using `finite J` `finite I` `I \<inter> J = {}` by (simp add: setprod_Un_one) |
|
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
|
776 |
also have "\<dots> = IJ.\<mu> A" |
41661 | 777 |
using `finite J` `finite I` F unfolding A |
778 |
by (intro IJ.measure_times[symmetric]) auto |
|
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41705
diff
changeset
|
779 |
finally show "IJ.\<mu> A = P.\<mu> (?X A)" by (rule sym) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
780 |
qed (rule measurable_merge[OF IJ]) |
41661 | 781 |
qed |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
782 |
|
41831 | 783 |
lemma (in product_sigma_finite) measure_preserving_merge: |
784 |
assumes IJ: "I \<inter> J = {}" and "finite I" "finite J" |
|
785 |
shows "(\<lambda>(x,y). merge I x J y) \<in> measure_preserving (Pi\<^isub>M I M \<Otimes>\<^isub>M Pi\<^isub>M J M) (Pi\<^isub>M (I \<union> J) M)" |
|
786 |
by (intro measure_preservingI measurable_merge[OF IJ] measure_fold[symmetric] assms) |
|
787 |
||
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
788 |
lemma (in product_sigma_finite) product_positive_integral_fold: |
41831 | 789 |
assumes IJ[simp]: "I \<inter> J = {}" "finite I" "finite J" |
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
|
790 |
and f: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)" |
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
|
791 |
shows "integral\<^isup>P (Pi\<^isub>M (I \<union> J) M) f = |
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
|
792 |
(\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (merge I x J y) \<partial>(Pi\<^isub>M J M)) \<partial>(Pi\<^isub>M I M))" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
793 |
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
|
794 |
interpret I: finite_product_sigma_finite M I by default fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
795 |
interpret J: finite_product_sigma_finite M J by default fact |
41831 | 796 |
interpret P: pair_sigma_finite "Pi\<^isub>M I M" "Pi\<^isub>M J M" by default |
797 |
interpret IJ: finite_product_sigma_finite M "I \<union> J" by default simp |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
798 |
have P_borel: "(\<lambda>x. f (case x of (x, y) \<Rightarrow> merge I x J y)) \<in> borel_measurable P.P" |
41831 | 799 |
using measurable_comp[OF measurable_merge[OF IJ(1)] f] by (simp add: comp_def) |
41661 | 800 |
show ?thesis |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
801 |
unfolding P.positive_integral_fst_measurable[OF P_borel, simplified] |
41661 | 802 |
proof (rule P.positive_integral_vimage) |
803 |
show "sigma_algebra IJ.P" by default |
|
41831 | 804 |
show "(\<lambda>(x, y). merge I x J y) \<in> measure_preserving P.P IJ.P" |
805 |
using IJ by (rule measure_preserving_merge) |
|
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
|
806 |
show "f \<in> borel_measurable IJ.P" using f by simp |
41661 | 807 |
qed |
40859 | 808 |
qed |
809 |
||
41831 | 810 |
lemma (in product_sigma_finite) measure_preserving_component_singelton: |
811 |
"(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)" |
|
812 |
proof (intro measure_preservingI measurable_component_singleton) |
|
813 |
interpret I: finite_product_sigma_finite M "{i}" by default simp |
|
814 |
fix A let ?P = "(\<lambda>x. x i) -` A \<inter> space I.P" |
|
815 |
assume A: "A \<in> sets (M i)" |
|
816 |
then have *: "?P = {i} \<rightarrow>\<^isub>E A" using sets_into_space by auto |
|
817 |
show "I.\<mu> ?P = M.\<mu> i A" unfolding * |
|
818 |
using A I.measure_times[of "\<lambda>_. A"] by auto |
|
819 |
qed auto |
|
820 |
||
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
821 |
lemma (in product_sigma_finite) product_positive_integral_singleton: |
40859 | 822 |
assumes f: "f \<in> borel_measurable (M i)" |
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
|
823 |
shows "integral\<^isup>P (Pi\<^isub>M {i} M) (\<lambda>x. f (x i)) = integral\<^isup>P (M i) f" |
40859 | 824 |
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
|
825 |
interpret I: finite_product_sigma_finite M "{i}" by default simp |
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
|
826 |
show ?thesis |
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
|
827 |
proof (rule I.positive_integral_vimage[symmetric]) |
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
|
828 |
show "sigma_algebra (M i)" by (rule sigma_algebras) |
41831 | 829 |
show "(\<lambda>x. x i) \<in> measure_preserving (Pi\<^isub>M {i} M) (M i)" |
830 |
by (rule measure_preserving_component_singelton) |
|
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
|
831 |
show "f \<in> borel_measurable (M i)" by fact |
41661 | 832 |
qed |
40859 | 833 |
qed |
834 |
||
41096 | 835 |
lemma (in product_sigma_finite) product_positive_integral_insert: |
836 |
assumes [simp]: "finite I" "i \<notin> I" |
|
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
|
837 |
and f: "f \<in> borel_measurable (Pi\<^isub>M (insert i I) M)" |
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
|
838 |
shows "integral\<^isup>P (Pi\<^isub>M (insert i I) M) f = (\<integral>\<^isup>+ x. (\<integral>\<^isup>+ y. f (x(i := y)) \<partial>(M i)) \<partial>(Pi\<^isub>M I M))" |
41096 | 839 |
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
|
840 |
interpret I: finite_product_sigma_finite M I by default auto |
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
|
841 |
interpret i: finite_product_sigma_finite M "{i}" by default auto |
41096 | 842 |
interpret P: pair_sigma_algebra I.P i.P .. |
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
|
843 |
have IJ: "I \<inter> {i} = {}" and insert: "I \<union> {i} = insert i I" |
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
|
844 |
using f by auto |
41096 | 845 |
show ?thesis |
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
|
846 |
unfolding product_positive_integral_fold[OF IJ, unfolded insert, simplified, OF f] |
41096 | 847 |
proof (rule I.positive_integral_cong, subst product_positive_integral_singleton) |
848 |
fix x assume x: "x \<in> space I.P" |
|
849 |
let "?f y" = "f (restrict (x(i := y)) (insert i I))" |
|
850 |
have f'_eq: "\<And>y. ?f y = f (x(i := y))" |
|
851 |
using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
852 |
show "?f \<in> borel_measurable (M i)" unfolding f'_eq |
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
|
853 |
using measurable_comp[OF measurable_component_update f] x `i \<notin> I` |
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
|
854 |
by (simp add: comp_def) |
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
|
855 |
show "integral\<^isup>P (M i) ?f = \<integral>\<^isup>+ y. f (x(i:=y)) \<partial>M i" |
41096 | 856 |
unfolding f'_eq by simp |
857 |
qed |
|
858 |
qed |
|
859 |
||
860 |
lemma (in product_sigma_finite) product_positive_integral_setprod: |
|
43920 | 861 |
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal" |
41096 | 862 |
assumes "finite I" and borel: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
863 |
and pos: "\<And>i x. i \<in> I \<Longrightarrow> 0 \<le> f i x" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
864 |
shows "(\<integral>\<^isup>+ x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>P (M i) (f i))" |
41096 | 865 |
using assms proof induct |
866 |
case empty |
|
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
|
867 |
interpret finite_product_sigma_finite M "{}" by default auto |
41096 | 868 |
then show ?case by simp |
869 |
next |
|
870 |
case (insert i I) |
|
871 |
note `finite I`[intro, simp] |
|
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
|
872 |
interpret I: finite_product_sigma_finite M I by default auto |
41096 | 873 |
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))" |
874 |
using insert by (auto intro!: setprod_cong) |
|
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
|
875 |
have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> (\<lambda>x. (\<Prod>i\<in>J. f i (x i))) \<in> borel_measurable (Pi\<^isub>M J M)" |
41096 | 876 |
using sets_into_space insert |
43920 | 877 |
by (intro sigma_algebra.borel_measurable_ereal_setprod sigma_algebra_product_algebra |
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
|
878 |
measurable_comp[OF measurable_component_singleton, unfolded comp_def]) |
41096 | 879 |
auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
880 |
then show ?case |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
881 |
apply (simp add: product_positive_integral_insert[OF insert(1,2) prod]) |
43920 | 882 |
apply (simp add: insert * pos borel setprod_ereal_pos M.positive_integral_multc) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
883 |
apply (subst I.positive_integral_cmult) |
43920 | 884 |
apply (auto simp add: pos borel insert setprod_ereal_pos positive_integral_positive) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
885 |
done |
41096 | 886 |
qed |
887 |
||
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
888 |
lemma (in product_sigma_finite) product_integral_singleton: |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
889 |
assumes f: "f \<in> borel_measurable (M i)" |
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
|
890 |
shows "(\<integral>x. f (x i) \<partial>Pi\<^isub>M {i} M) = integral\<^isup>L (M i) f" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
891 |
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
|
892 |
interpret I: finite_product_sigma_finite M "{i}" by default simp |
43920 | 893 |
have *: "(\<lambda>x. ereal (f x)) \<in> borel_measurable (M i)" |
894 |
"(\<lambda>x. ereal (- f x)) \<in> borel_measurable (M i)" |
|
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
895 |
using assms by auto |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
896 |
show ?thesis |
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
|
897 |
unfolding lebesgue_integral_def *[THEN product_positive_integral_singleton] .. |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
898 |
qed |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
899 |
|
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
900 |
lemma (in product_sigma_finite) product_integral_fold: |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
901 |
assumes IJ[simp]: "I \<inter> J = {}" and fin: "finite I" "finite J" |
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
|
902 |
and f: "integrable (Pi\<^isub>M (I \<union> J) M) f" |
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
|
903 |
shows "integral\<^isup>L (Pi\<^isub>M (I \<union> J) M) f = (\<integral>x. (\<integral>y. f (merge I x J y) \<partial>Pi\<^isub>M J M) \<partial>Pi\<^isub>M I M)" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
904 |
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
|
905 |
interpret I: finite_product_sigma_finite M I by default fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
906 |
interpret J: finite_product_sigma_finite M J by default fact |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
907 |
have "finite (I \<union> J)" using fin by auto |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
908 |
interpret IJ: finite_product_sigma_finite M "I \<union> J" by default fact |
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
909 |
interpret P: pair_sigma_finite I.P J.P by default |
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
|
910 |
let ?M = "\<lambda>(x, y). merge I x J y" |
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
|
911 |
let ?f = "\<lambda>x. f (?M x)" |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
912 |
show ?thesis |
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
|
913 |
proof (subst P.integrable_fst_measurable(2)[of ?f, simplified]) |
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
|
914 |
have 1: "sigma_algebra IJ.P" by default |
41831 | 915 |
have 2: "?M \<in> measure_preserving P.P IJ.P" using measure_preserving_merge[OF assms(1,2,3)] . |
916 |
have 3: "integrable (Pi\<^isub>M (I \<union> J) M) f" by fact |
|
917 |
then have 4: "f \<in> borel_measurable (Pi\<^isub>M (I \<union> J) M)" |
|
918 |
by (simp add: integrable_def) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
919 |
show "integrable P.P ?f" |
41831 | 920 |
by (rule P.integrable_vimage[where f=f, OF 1 2 3]) |
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
|
921 |
show "integral\<^isup>L IJ.P f = integral\<^isup>L P.P ?f" |
41831 | 922 |
by (rule P.integral_vimage[where f=f, OF 1 2 4]) |
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
|
923 |
qed |
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
924 |
qed |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
925 |
|
41096 | 926 |
lemma (in product_sigma_finite) product_integral_insert: |
927 |
assumes [simp]: "finite I" "i \<notin> I" |
|
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
|
928 |
and f: "integrable (Pi\<^isub>M (insert i I) M) f" |
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
|
929 |
shows "integral\<^isup>L (Pi\<^isub>M (insert i I) M) f = (\<integral>x. (\<integral>y. f (x(i:=y)) \<partial>M i) \<partial>Pi\<^isub>M I M)" |
41096 | 930 |
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
|
931 |
interpret I: finite_product_sigma_finite M I by default auto |
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
|
932 |
interpret I': finite_product_sigma_finite M "insert i I" by default auto |
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
|
933 |
interpret i: finite_product_sigma_finite M "{i}" by default auto |
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
|
934 |
interpret P: pair_sigma_finite I.P i.P .. |
41096 | 935 |
have IJ: "I \<inter> {i} = {}" by auto |
936 |
show ?thesis |
|
937 |
unfolding product_integral_fold[OF IJ, simplified, OF f] |
|
938 |
proof (rule I.integral_cong, subst product_integral_singleton) |
|
939 |
fix x assume x: "x \<in> space I.P" |
|
940 |
let "?f y" = "f (restrict (x(i := y)) (insert i I))" |
|
941 |
have f'_eq: "\<And>y. ?f y = f (x(i := y))" |
|
942 |
using x by (auto intro!: arg_cong[where f=f] simp: fun_eq_iff extensional_def) |
|
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
943 |
have f: "f \<in> borel_measurable I'.P" using f unfolding integrable_def by auto |
41096 | 944 |
show "?f \<in> borel_measurable (M i)" |
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
|
945 |
unfolding measurable_cong[OF f'_eq] |
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
|
946 |
using measurable_comp[OF measurable_component_update f] x `i \<notin> I` |
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
|
947 |
by (simp add: comp_def) |
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
|
948 |
show "integral\<^isup>L (M i) ?f = integral\<^isup>L (M i) (\<lambda>y. f (x(i := y)))" |
41096 | 949 |
unfolding f'_eq by simp |
950 |
qed |
|
951 |
qed |
|
952 |
||
953 |
lemma (in product_sigma_finite) product_integrable_setprod: |
|
954 |
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real" |
|
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
|
955 |
assumes [simp]: "finite I" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)" |
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
|
956 |
shows "integrable (Pi\<^isub>M I M) (\<lambda>x. (\<Prod>i\<in>I. f i (x i)))" (is "integrable _ ?f") |
41096 | 957 |
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
|
958 |
interpret finite_product_sigma_finite M I by default fact |
41096 | 959 |
have f: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable (M i)" |
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
|
960 |
using integrable unfolding integrable_def by auto |
41096 | 961 |
then have borel: "?f \<in> borel_measurable P" |
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
|
962 |
using measurable_comp[OF measurable_component_singleton f] |
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
|
963 |
by (auto intro!: borel_measurable_setprod simp: comp_def) |
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
|
964 |
moreover have "integrable (Pi\<^isub>M I M) (\<lambda>x. \<bar>\<Prod>i\<in>I. f i (x i)\<bar>)" |
41096 | 965 |
proof (unfold integrable_def, intro conjI) |
966 |
show "(\<lambda>x. abs (?f x)) \<in> borel_measurable P" |
|
967 |
using borel by auto |
|
43920 | 968 |
have "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. (\<Prod>i\<in>I. ereal (abs (f i (x i)))) \<partial>P)" |
969 |
by (simp add: setprod_ereal abs_setprod) |
|
970 |
also have "\<dots> = (\<Prod>i\<in>I. (\<integral>\<^isup>+x. ereal (abs (f i x)) \<partial>M i))" |
|
41096 | 971 |
using f by (subst product_positive_integral_setprod) auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
972 |
also have "\<dots> < \<infinity>" |
41096 | 973 |
using integrable[THEN M.integrable_abs] |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
974 |
by (simp add: setprod_PInf integrable_def M.positive_integral_positive) |
43920 | 975 |
finally show "(\<integral>\<^isup>+x. ereal (abs (?f x)) \<partial>P) \<noteq> \<infinity>" by auto |
976 |
have "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) = (\<integral>\<^isup>+x. 0 \<partial>P)" |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
977 |
by (intro positive_integral_cong_pos) auto |
43920 | 978 |
then show "(\<integral>\<^isup>+x. ereal (- abs (?f x)) \<partial>P) \<noteq> \<infinity>" by simp |
41096 | 979 |
qed |
980 |
ultimately show ?thesis |
|
981 |
by (rule integrable_abs_iff[THEN iffD1]) |
|
982 |
qed |
|
983 |
||
984 |
lemma (in product_sigma_finite) product_integral_setprod: |
|
985 |
fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> real" |
|
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
|
986 |
assumes "finite I" "I \<noteq> {}" and integrable: "\<And>i. i \<in> I \<Longrightarrow> integrable (M i) (f i)" |
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
|
987 |
shows "(\<integral>x. (\<Prod>i\<in>I. f i (x i)) \<partial>Pi\<^isub>M I M) = (\<Prod>i\<in>I. integral\<^isup>L (M i) (f i))" |
41096 | 988 |
using assms proof (induct rule: finite_ne_induct) |
989 |
case (singleton i) |
|
990 |
then show ?case by (simp add: product_integral_singleton integrable_def) |
|
991 |
next |
|
992 |
case (insert i I) |
|
993 |
then have iI: "finite (insert i I)" by auto |
|
994 |
then have prod: "\<And>J. J \<subseteq> insert i I \<Longrightarrow> |
|
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
|
995 |
integrable (Pi\<^isub>M J M) (\<lambda>x. (\<Prod>i\<in>J. f i (x i)))" |
41096 | 996 |
by (intro product_integrable_setprod insert(5)) (auto intro: finite_subset) |
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
|
997 |
interpret I: finite_product_sigma_finite M I by default fact |
41096 | 998 |
have *: "\<And>x y. (\<Prod>j\<in>I. f j (if j = i then y else x j)) = (\<Prod>j\<in>I. f j (x j))" |
999 |
using `i \<notin> I` by (auto intro!: setprod_cong) |
|
1000 |
show ?case |
|
1001 |
unfolding product_integral_insert[OF insert(1,3) prod[OF subset_refl]] |
|
1002 |
by (simp add: * insert integral_multc I.integral_cmult[OF prod] subset_insertI) |
|
1003 |
qed |
|
1004 |
||
40859 | 1005 |
end |