author | hoelzl |
Mon, 06 Dec 2010 19:54:56 +0100 | |
changeset 41026 | bea75746dc9d |
parent 41023 | 9118eb4eb8dc |
child 41095 | c335d880ff82 |
permissions | -rw-r--r-- |
38656 | 1 |
(* Author: Armin Heller, Johannes Hoelzl, TU Muenchen *) |
2 |
||
35582 | 3 |
header {*Lebesgue Integration*} |
4 |
||
38656 | 5 |
theory Lebesgue_Integration |
40859 | 6 |
imports Measure Borel_Space |
35582 | 7 |
begin |
8 |
||
38656 | 9 |
lemma sums_If_finite: |
10 |
assumes finite: "finite {r. P r}" |
|
11 |
shows "(\<lambda>r. if P r then f r else 0) sums (\<Sum>r\<in>{r. P r}. f r)" (is "?F sums _") |
|
12 |
proof cases |
|
13 |
assume "{r. P r} = {}" hence "\<And>r. \<not> P r" by auto |
|
14 |
thus ?thesis by (simp add: sums_zero) |
|
15 |
next |
|
16 |
assume not_empty: "{r. P r} \<noteq> {}" |
|
17 |
have "?F sums (\<Sum>r = 0..< Suc (Max {r. P r}). ?F r)" |
|
18 |
by (rule series_zero) |
|
19 |
(auto simp add: Max_less_iff[OF finite not_empty] less_eq_Suc_le[symmetric]) |
|
20 |
also have "(\<Sum>r = 0..< Suc (Max {r. P r}). ?F r) = (\<Sum>r\<in>{r. P r}. f r)" |
|
21 |
by (subst setsum_cases) |
|
22 |
(auto intro!: setsum_cong simp: Max_ge_iff[OF finite not_empty] less_Suc_eq_le) |
|
23 |
finally show ?thesis . |
|
24 |
qed |
|
35582 | 25 |
|
38656 | 26 |
lemma sums_single: |
27 |
"(\<lambda>r. if r = i then f r else 0) sums f i" |
|
28 |
using sums_If_finite[of "\<lambda>r. r = i" f] by simp |
|
29 |
||
30 |
section "Simple function" |
|
35582 | 31 |
|
38656 | 32 |
text {* |
33 |
||
34 |
Our simple functions are not restricted to positive real numbers. Instead |
|
35 |
they are just functions with a finite range and are measurable when singleton |
|
36 |
sets are measurable. |
|
35582 | 37 |
|
38656 | 38 |
*} |
39 |
||
40 |
definition (in sigma_algebra) "simple_function g \<longleftrightarrow> |
|
41 |
finite (g ` space M) \<and> |
|
42 |
(\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)" |
|
36624 | 43 |
|
38656 | 44 |
lemma (in sigma_algebra) simple_functionD: |
45 |
assumes "simple_function g" |
|
40875 | 46 |
shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M" |
40871 | 47 |
proof - |
48 |
show "finite (g ` space M)" |
|
49 |
using assms unfolding simple_function_def by auto |
|
40875 | 50 |
have "g -` X \<inter> space M = g -` (X \<inter> g`space M) \<inter> space M" by auto |
51 |
also have "\<dots> = (\<Union>x\<in>X \<inter> g`space M. g-`{x} \<inter> space M)" by auto |
|
52 |
finally show "g -` X \<inter> space M \<in> sets M" using assms |
|
53 |
by (auto intro!: finite_UN simp del: UN_simps simp: simple_function_def) |
|
40871 | 54 |
qed |
36624 | 55 |
|
38656 | 56 |
lemma (in sigma_algebra) simple_function_indicator_representation: |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
57 |
fixes f ::"'a \<Rightarrow> pextreal" |
38656 | 58 |
assumes f: "simple_function f" and x: "x \<in> space M" |
59 |
shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)" |
|
60 |
(is "?l = ?r") |
|
61 |
proof - |
|
38705 | 62 |
have "?r = (\<Sum>y \<in> f ` space M. |
38656 | 63 |
(if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))" |
64 |
by (auto intro!: setsum_cong2) |
|
65 |
also have "... = f x * indicator (f -` {f x} \<inter> space M) x" |
|
66 |
using assms by (auto dest: simple_functionD simp: setsum_delta) |
|
67 |
also have "... = f x" using x by (auto simp: indicator_def) |
|
68 |
finally show ?thesis by auto |
|
69 |
qed |
|
36624 | 70 |
|
38656 | 71 |
lemma (in measure_space) simple_function_notspace: |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
72 |
"simple_function (\<lambda>x. h x * indicator (- space M) x::pextreal)" (is "simple_function ?h") |
35692 | 73 |
proof - |
38656 | 74 |
have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto |
75 |
hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) |
|
76 |
have "?h -` {0} \<inter> space M = space M" by auto |
|
77 |
thus ?thesis unfolding simple_function_def by auto |
|
78 |
qed |
|
79 |
||
80 |
lemma (in sigma_algebra) simple_function_cong: |
|
81 |
assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t" |
|
82 |
shows "simple_function f \<longleftrightarrow> simple_function g" |
|
83 |
proof - |
|
84 |
have "f ` space M = g ` space M" |
|
85 |
"\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M" |
|
86 |
using assms by (auto intro!: image_eqI) |
|
87 |
thus ?thesis unfolding simple_function_def using assms by simp |
|
88 |
qed |
|
89 |
||
90 |
lemma (in sigma_algebra) borel_measurable_simple_function: |
|
91 |
assumes "simple_function f" |
|
92 |
shows "f \<in> borel_measurable M" |
|
93 |
proof (rule borel_measurableI) |
|
94 |
fix S |
|
95 |
let ?I = "f ` (f -` S \<inter> space M)" |
|
96 |
have *: "(\<Union>x\<in>?I. f -` {x} \<inter> space M) = f -` S \<inter> space M" (is "?U = _") by auto |
|
97 |
have "finite ?I" |
|
98 |
using assms unfolding simple_function_def by (auto intro: finite_subset) |
|
99 |
hence "?U \<in> sets M" |
|
100 |
apply (rule finite_UN) |
|
101 |
using assms unfolding simple_function_def by auto |
|
102 |
thus "f -` S \<inter> space M \<in> sets M" unfolding * . |
|
35692 | 103 |
qed |
104 |
||
38656 | 105 |
lemma (in sigma_algebra) simple_function_borel_measurable: |
106 |
fixes f :: "'a \<Rightarrow> 'x::t2_space" |
|
107 |
assumes "f \<in> borel_measurable M" and "finite (f ` space M)" |
|
108 |
shows "simple_function f" |
|
109 |
using assms unfolding simple_function_def |
|
110 |
by (auto intro: borel_measurable_vimage) |
|
111 |
||
112 |
lemma (in sigma_algebra) simple_function_const[intro, simp]: |
|
113 |
"simple_function (\<lambda>x. c)" |
|
114 |
by (auto intro: finite_subset simp: simple_function_def) |
|
115 |
||
116 |
lemma (in sigma_algebra) simple_function_compose[intro, simp]: |
|
117 |
assumes "simple_function f" |
|
118 |
shows "simple_function (g \<circ> f)" |
|
119 |
unfolding simple_function_def |
|
120 |
proof safe |
|
121 |
show "finite ((g \<circ> f) ` space M)" |
|
122 |
using assms unfolding simple_function_def by (auto simp: image_compose) |
|
123 |
next |
|
124 |
fix x assume "x \<in> space M" |
|
125 |
let ?G = "g -` {g (f x)} \<inter> (f`space M)" |
|
126 |
have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M = |
|
127 |
(\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto |
|
128 |
show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M" |
|
129 |
using assms unfolding simple_function_def * |
|
130 |
by (rule_tac finite_UN) (auto intro!: finite_UN) |
|
131 |
qed |
|
132 |
||
133 |
lemma (in sigma_algebra) simple_function_indicator[intro, simp]: |
|
134 |
assumes "A \<in> sets M" |
|
135 |
shows "simple_function (indicator A)" |
|
35692 | 136 |
proof - |
38656 | 137 |
have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _") |
138 |
by (auto simp: indicator_def) |
|
139 |
hence "finite ?S" by (rule finite_subset) simp |
|
140 |
moreover have "- A \<inter> space M = space M - A" by auto |
|
141 |
ultimately show ?thesis unfolding simple_function_def |
|
142 |
using assms by (auto simp: indicator_def_raw) |
|
35692 | 143 |
qed |
144 |
||
38656 | 145 |
lemma (in sigma_algebra) simple_function_Pair[intro, simp]: |
146 |
assumes "simple_function f" |
|
147 |
assumes "simple_function g" |
|
148 |
shows "simple_function (\<lambda>x. (f x, g x))" (is "simple_function ?p") |
|
149 |
unfolding simple_function_def |
|
150 |
proof safe |
|
151 |
show "finite (?p ` space M)" |
|
152 |
using assms unfolding simple_function_def |
|
153 |
by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto |
|
154 |
next |
|
155 |
fix x assume "x \<in> space M" |
|
156 |
have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M = |
|
157 |
(f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)" |
|
158 |
by auto |
|
159 |
with `x \<in> space M` show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M" |
|
160 |
using assms unfolding simple_function_def by auto |
|
161 |
qed |
|
35692 | 162 |
|
38656 | 163 |
lemma (in sigma_algebra) simple_function_compose1: |
164 |
assumes "simple_function f" |
|
165 |
shows "simple_function (\<lambda>x. g (f x))" |
|
166 |
using simple_function_compose[OF assms, of g] |
|
167 |
by (simp add: comp_def) |
|
35582 | 168 |
|
38656 | 169 |
lemma (in sigma_algebra) simple_function_compose2: |
170 |
assumes "simple_function f" and "simple_function g" |
|
171 |
shows "simple_function (\<lambda>x. h (f x) (g x))" |
|
172 |
proof - |
|
173 |
have "simple_function ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))" |
|
174 |
using assms by auto |
|
175 |
thus ?thesis by (simp_all add: comp_def) |
|
176 |
qed |
|
35582 | 177 |
|
38656 | 178 |
lemmas (in sigma_algebra) simple_function_add[intro, simp] = simple_function_compose2[where h="op +"] |
179 |
and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"] |
|
180 |
and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] |
|
181 |
and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"] |
|
182 |
and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"] |
|
183 |
and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] |
|
184 |
||
185 |
lemma (in sigma_algebra) simple_function_setsum[intro, simp]: |
|
186 |
assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function (f i)" |
|
187 |
shows "simple_function (\<lambda>x. \<Sum>i\<in>P. f i x)" |
|
188 |
proof cases |
|
189 |
assume "finite P" from this assms show ?thesis by induct auto |
|
190 |
qed auto |
|
35582 | 191 |
|
38656 | 192 |
lemma (in sigma_algebra) simple_function_le_measurable: |
193 |
assumes "simple_function f" "simple_function g" |
|
194 |
shows "{x \<in> space M. f x \<le> g x} \<in> sets M" |
|
195 |
proof - |
|
196 |
have *: "{x \<in> space M. f x \<le> g x} = |
|
197 |
(\<Union>(F, G)\<in>f`space M \<times> g`space M. |
|
198 |
if F \<le> G then (f -` {F} \<inter> space M) \<inter> (g -` {G} \<inter> space M) else {})" |
|
199 |
apply (auto split: split_if_asm) |
|
200 |
apply (rule_tac x=x in bexI) |
|
201 |
apply (rule_tac x=x in bexI) |
|
202 |
by simp_all |
|
203 |
have **: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow> |
|
204 |
(f -` {f x} \<inter> space M) \<inter> (g -` {g y} \<inter> space M) \<in> sets M" |
|
205 |
using assms unfolding simple_function_def by auto |
|
206 |
have "finite (f`space M \<times> g`space M)" |
|
207 |
using assms unfolding simple_function_def by auto |
|
208 |
thus ?thesis unfolding * |
|
209 |
apply (rule finite_UN) |
|
210 |
using assms unfolding simple_function_def |
|
211 |
by (auto intro!: **) |
|
35582 | 212 |
qed |
213 |
||
38656 | 214 |
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence: |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
215 |
fixes u :: "'a \<Rightarrow> pextreal" |
38656 | 216 |
assumes u: "u \<in> borel_measurable M" |
217 |
shows "\<exists>f. (\<forall>i. simple_function (f i) \<and> (\<forall>x\<in>space M. f i x \<noteq> \<omega>)) \<and> f \<up> u" |
|
35582 | 218 |
proof - |
38656 | 219 |
have "\<exists>f. \<forall>x j. (of_nat j \<le> u x \<longrightarrow> f x j = j*2^j) \<and> |
220 |
(u x < of_nat j \<longrightarrow> of_nat (f x j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f x j)))" |
|
221 |
(is "\<exists>f. \<forall>x j. ?P x j (f x j)") |
|
222 |
proof(rule choice, rule, rule choice, rule) |
|
223 |
fix x j show "\<exists>n. ?P x j n" |
|
224 |
proof cases |
|
225 |
assume *: "u x < of_nat j" |
|
226 |
then obtain r where r: "u x = Real r" "0 \<le> r" by (cases "u x") auto |
|
227 |
from reals_Archimedean6a[of "r * 2^j"] |
|
228 |
obtain n :: nat where "real n \<le> r * 2 ^ j" "r * 2 ^ j < real (Suc n)" |
|
229 |
using `0 \<le> r` by (auto simp: zero_le_power zero_le_mult_iff) |
|
230 |
thus ?thesis using r * by (auto intro!: exI[of _ n]) |
|
231 |
qed auto |
|
35582 | 232 |
qed |
38656 | 233 |
then obtain f where top: "\<And>j x. of_nat j \<le> u x \<Longrightarrow> f x j = j*2^j" and |
234 |
upper: "\<And>j x. u x < of_nat j \<Longrightarrow> of_nat (f x j) \<le> u x * 2^j" and |
|
235 |
lower: "\<And>j x. u x < of_nat j \<Longrightarrow> u x * 2^j < of_nat (Suc (f x j))" by blast |
|
35582 | 236 |
|
38656 | 237 |
{ fix j x P |
238 |
assume 1: "of_nat j \<le> u x \<Longrightarrow> P (j * 2^j)" |
|
239 |
assume 2: "\<And>k. \<lbrakk> u x < of_nat j ; of_nat k \<le> u x * 2^j ; u x * 2^j < of_nat (Suc k) \<rbrakk> \<Longrightarrow> P k" |
|
240 |
have "P (f x j)" |
|
241 |
proof cases |
|
242 |
assume "of_nat j \<le> u x" thus "P (f x j)" |
|
243 |
using top[of j x] 1 by auto |
|
244 |
next |
|
245 |
assume "\<not> of_nat j \<le> u x" |
|
246 |
hence "u x < of_nat j" "of_nat (f x j) \<le> u x * 2^j" "u x * 2^j < of_nat (Suc (f x j))" |
|
247 |
using upper lower by auto |
|
248 |
from 2[OF this] show "P (f x j)" . |
|
249 |
qed } |
|
250 |
note fI = this |
|
251 |
||
252 |
{ fix j x |
|
253 |
have "f x j = j * 2 ^ j \<longleftrightarrow> of_nat j \<le> u x" |
|
254 |
by (rule fI, simp, cases "u x") (auto split: split_if_asm) } |
|
255 |
note f_eq = this |
|
35582 | 256 |
|
38656 | 257 |
{ fix j x |
258 |
have "f x j \<le> j * 2 ^ j" |
|
259 |
proof (rule fI) |
|
260 |
fix k assume *: "u x < of_nat j" |
|
261 |
assume "of_nat k \<le> u x * 2 ^ j" |
|
262 |
also have "\<dots> \<le> of_nat (j * 2^j)" |
|
263 |
using * by (cases "u x") (auto simp: zero_le_mult_iff) |
|
264 |
finally show "k \<le> j*2^j" by (auto simp del: real_of_nat_mult) |
|
265 |
qed simp } |
|
266 |
note f_upper = this |
|
35582 | 267 |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
268 |
let "?g j x" = "of_nat (f x j) / 2^j :: pextreal" |
38656 | 269 |
show ?thesis unfolding simple_function_def isoton_fun_expand unfolding isoton_def |
270 |
proof (safe intro!: exI[of _ ?g]) |
|
271 |
fix j |
|
272 |
have *: "?g j ` space M \<subseteq> (\<lambda>x. of_nat x / 2^j) ` {..j * 2^j}" |
|
273 |
using f_upper by auto |
|
274 |
thus "finite (?g j ` space M)" by (rule finite_subset) auto |
|
275 |
next |
|
276 |
fix j t assume "t \<in> space M" |
|
277 |
have **: "?g j -` {?g j t} \<inter> space M = {x \<in> space M. f t j = f x j}" |
|
278 |
by (auto simp: power_le_zero_eq Real_eq_Real mult_le_0_iff) |
|
35582 | 279 |
|
38656 | 280 |
show "?g j -` {?g j t} \<inter> space M \<in> sets M" |
281 |
proof cases |
|
282 |
assume "of_nat j \<le> u t" |
|
283 |
hence "?g j -` {?g j t} \<inter> space M = {x\<in>space M. of_nat j \<le> u x}" |
|
284 |
unfolding ** f_eq[symmetric] by auto |
|
285 |
thus "?g j -` {?g j t} \<inter> space M \<in> sets M" |
|
286 |
using u by auto |
|
35582 | 287 |
next |
38656 | 288 |
assume not_t: "\<not> of_nat j \<le> u t" |
289 |
hence less: "f t j < j*2^j" using f_eq[symmetric] `f t j \<le> j*2^j` by auto |
|
290 |
have split_vimage: "?g j -` {?g j t} \<inter> space M = |
|
291 |
{x\<in>space M. of_nat (f t j)/2^j \<le> u x} \<inter> {x\<in>space M. u x < of_nat (Suc (f t j))/2^j}" |
|
292 |
unfolding ** |
|
293 |
proof safe |
|
294 |
fix x assume [simp]: "f t j = f x j" |
|
295 |
have *: "\<not> of_nat j \<le> u x" using not_t f_eq[symmetric] by simp |
|
296 |
hence "of_nat (f t j) \<le> u x * 2^j \<and> u x * 2^j < of_nat (Suc (f t j))" |
|
297 |
using upper lower by auto |
|
298 |
hence "of_nat (f t j) / 2^j \<le> u x \<and> u x < of_nat (Suc (f t j))/2^j" using * |
|
299 |
by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) |
|
300 |
thus "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j" by auto |
|
301 |
next |
|
302 |
fix x |
|
303 |
assume "of_nat (f t j) / 2^j \<le> u x" "u x < of_nat (Suc (f t j))/2^j" |
|
304 |
hence "of_nat (f t j) \<le> u x * 2 ^ j \<and> u x * 2 ^ j < of_nat (Suc (f t j))" |
|
305 |
by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) |
|
306 |
hence 1: "of_nat (f t j) \<le> u x * 2 ^ j" and 2: "u x * 2 ^ j < of_nat (Suc (f t j))" by auto |
|
307 |
note 2 |
|
308 |
also have "\<dots> \<le> of_nat (j*2^j)" |
|
309 |
using less by (auto simp: zero_le_mult_iff simp del: real_of_nat_mult) |
|
310 |
finally have bound_ux: "u x < of_nat j" |
|
311 |
by (cases "u x") (auto simp: zero_le_mult_iff power_le_zero_eq) |
|
312 |
show "f t j = f x j" |
|
313 |
proof (rule antisym) |
|
314 |
from 1 lower[OF bound_ux] |
|
315 |
show "f t j \<le> f x j" by (cases "u x") (auto split: split_if_asm) |
|
316 |
from upper[OF bound_ux] 2 |
|
317 |
show "f x j \<le> f t j" by (cases "u x") (auto split: split_if_asm) |
|
318 |
qed |
|
319 |
qed |
|
320 |
show ?thesis unfolding split_vimage using u by auto |
|
35582 | 321 |
qed |
38656 | 322 |
next |
323 |
fix j t assume "?g j t = \<omega>" thus False by (auto simp: power_le_zero_eq) |
|
324 |
next |
|
325 |
fix t |
|
326 |
{ fix i |
|
327 |
have "f t i * 2 \<le> f t (Suc i)" |
|
328 |
proof (rule fI) |
|
329 |
assume "of_nat (Suc i) \<le> u t" |
|
330 |
hence "of_nat i \<le> u t" by (cases "u t") auto |
|
331 |
thus "f t i * 2 \<le> Suc i * 2 ^ Suc i" unfolding f_eq[symmetric] by simp |
|
332 |
next |
|
333 |
fix k |
|
334 |
assume *: "u t * 2 ^ Suc i < of_nat (Suc k)" |
|
335 |
show "f t i * 2 \<le> k" |
|
336 |
proof (rule fI) |
|
337 |
assume "of_nat i \<le> u t" |
|
338 |
hence "of_nat (i * 2^Suc i) \<le> u t * 2^Suc i" |
|
339 |
by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq) |
|
340 |
also have "\<dots> < of_nat (Suc k)" using * by auto |
|
341 |
finally show "i * 2 ^ i * 2 \<le> k" |
|
342 |
by (auto simp del: real_of_nat_mult) |
|
343 |
next |
|
344 |
fix j assume "of_nat j \<le> u t * 2 ^ i" |
|
345 |
with * show "j * 2 \<le> k" by (cases "u t") (auto simp: zero_le_mult_iff power_le_zero_eq) |
|
346 |
qed |
|
347 |
qed |
|
348 |
thus "?g i t \<le> ?g (Suc i) t" |
|
349 |
by (auto simp: zero_le_mult_iff power_le_zero_eq divide_real_def[symmetric] field_simps) } |
|
350 |
hence mono: "mono (\<lambda>i. ?g i t)" unfolding mono_iff_le_Suc by auto |
|
35582 | 351 |
|
38656 | 352 |
show "(SUP j. of_nat (f t j) / 2 ^ j) = u t" |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
353 |
proof (rule pextreal_SUPI) |
38656 | 354 |
fix j show "of_nat (f t j) / 2 ^ j \<le> u t" |
355 |
proof (rule fI) |
|
356 |
assume "of_nat j \<le> u t" thus "of_nat (j * 2 ^ j) / 2 ^ j \<le> u t" |
|
357 |
by (cases "u t") (auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps) |
|
358 |
next |
|
359 |
fix k assume "of_nat k \<le> u t * 2 ^ j" |
|
360 |
thus "of_nat k / 2 ^ j \<le> u t" |
|
361 |
by (cases "u t") |
|
362 |
(auto simp: power_le_zero_eq divide_real_def[symmetric] field_simps zero_le_mult_iff) |
|
363 |
qed |
|
364 |
next |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
365 |
fix y :: pextreal assume *: "\<And>j. j \<in> UNIV \<Longrightarrow> of_nat (f t j) / 2 ^ j \<le> y" |
38656 | 366 |
show "u t \<le> y" |
367 |
proof (cases "u t") |
|
368 |
case (preal r) |
|
369 |
show ?thesis |
|
370 |
proof (rule ccontr) |
|
371 |
assume "\<not> u t \<le> y" |
|
372 |
then obtain p where p: "y = Real p" "0 \<le> p" "p < r" using preal by (cases y) auto |
|
373 |
with LIMSEQ_inverse_realpow_zero[of 2, unfolded LIMSEQ_iff, rule_format, of "r - p"] |
|
374 |
obtain n where n: "\<And>N. n \<le> N \<Longrightarrow> inverse (2^N) < r - p" by auto |
|
375 |
let ?N = "max n (natfloor r + 1)" |
|
376 |
have "u t < of_nat ?N" "n \<le> ?N" |
|
377 |
using ge_natfloor_plus_one_imp_gt[of r n] preal |
|
38705 | 378 |
using real_natfloor_add_one_gt |
379 |
by (auto simp: max_def real_of_nat_Suc) |
|
38656 | 380 |
from lower[OF this(1)] |
381 |
have "r < (real (f t ?N) + 1) / 2 ^ ?N" unfolding less_divide_eq |
|
382 |
using preal by (auto simp add: power_le_zero_eq zero_le_mult_iff real_of_nat_Suc split: split_if_asm) |
|
383 |
hence "u t < of_nat (f t ?N) / 2 ^ ?N + 1 / 2 ^ ?N" |
|
384 |
using preal by (auto simp: field_simps divide_real_def[symmetric]) |
|
385 |
with n[OF `n \<le> ?N`] p preal *[of ?N] |
|
386 |
show False |
|
387 |
by (cases "f t ?N") (auto simp: power_le_zero_eq split: split_if_asm) |
|
388 |
qed |
|
389 |
next |
|
390 |
case infinite |
|
391 |
{ fix j have "f t j = j*2^j" using top[of j t] infinite by simp |
|
392 |
hence "of_nat j \<le> y" using *[of j] |
|
393 |
by (cases y) (auto simp: power_le_zero_eq zero_le_mult_iff field_simps) } |
|
394 |
note all_less_y = this |
|
395 |
show ?thesis unfolding infinite |
|
396 |
proof (rule ccontr) |
|
397 |
assume "\<not> \<omega> \<le> y" then obtain r where r: "y = Real r" "0 \<le> r" by (cases y) auto |
|
398 |
moreover obtain n ::nat where "r < real n" using ex_less_of_nat by (auto simp: real_eq_of_nat) |
|
399 |
with all_less_y[of n] r show False by auto |
|
400 |
qed |
|
401 |
qed |
|
402 |
qed |
|
35582 | 403 |
qed |
404 |
qed |
|
405 |
||
38656 | 406 |
lemma (in sigma_algebra) borel_measurable_implies_simple_function_sequence': |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
407 |
fixes u :: "'a \<Rightarrow> pextreal" |
38656 | 408 |
assumes "u \<in> borel_measurable M" |
409 |
obtains (x) f where "f \<up> u" "\<And>i. simple_function (f i)" "\<And>i. \<omega>\<notin>f i`space M" |
|
35582 | 410 |
proof - |
38656 | 411 |
from borel_measurable_implies_simple_function_sequence[OF assms] |
412 |
obtain f where x: "\<And>i. simple_function (f i)" "f \<up> u" |
|
413 |
and fin: "\<And>i. \<And>x. x\<in>space M \<Longrightarrow> f i x \<noteq> \<omega>" by auto |
|
414 |
{ fix i from fin[of _ i] have "\<omega> \<notin> f i`space M" by fastsimp } |
|
415 |
with x show thesis by (auto intro!: that[of f]) |
|
416 |
qed |
|
417 |
||
39092 | 418 |
lemma (in sigma_algebra) simple_function_eq_borel_measurable: |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
419 |
fixes f :: "'a \<Rightarrow> pextreal" |
39092 | 420 |
shows "simple_function f \<longleftrightarrow> |
421 |
finite (f`space M) \<and> f \<in> borel_measurable M" |
|
422 |
using simple_function_borel_measurable[of f] |
|
423 |
borel_measurable_simple_function[of f] |
|
424 |
by (fastsimp simp: simple_function_def) |
|
425 |
||
426 |
lemma (in measure_space) simple_function_restricted: |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
427 |
fixes f :: "'a \<Rightarrow> pextreal" assumes "A \<in> sets M" |
39092 | 428 |
shows "sigma_algebra.simple_function (restricted_space A) f \<longleftrightarrow> simple_function (\<lambda>x. f x * indicator A x)" |
429 |
(is "sigma_algebra.simple_function ?R f \<longleftrightarrow> simple_function ?f") |
|
430 |
proof - |
|
431 |
interpret R: sigma_algebra ?R by (rule restricted_sigma_algebra[OF `A \<in> sets M`]) |
|
432 |
have "finite (f`A) \<longleftrightarrow> finite (?f`space M)" |
|
433 |
proof cases |
|
434 |
assume "A = space M" |
|
435 |
then have "f`A = ?f`space M" by (fastsimp simp: image_iff) |
|
436 |
then show ?thesis by simp |
|
437 |
next |
|
438 |
assume "A \<noteq> space M" |
|
439 |
then obtain x where x: "x \<in> space M" "x \<notin> A" |
|
440 |
using sets_into_space `A \<in> sets M` by auto |
|
441 |
have *: "?f`space M = f`A \<union> {0}" |
|
442 |
proof (auto simp add: image_iff) |
|
443 |
show "\<exists>x\<in>space M. f x = 0 \<or> indicator A x = 0" |
|
444 |
using x by (auto intro!: bexI[of _ x]) |
|
445 |
next |
|
446 |
fix x assume "x \<in> A" |
|
447 |
then show "\<exists>y\<in>space M. f x = f y * indicator A y" |
|
448 |
using `A \<in> sets M` sets_into_space by (auto intro!: bexI[of _ x]) |
|
449 |
next |
|
450 |
fix x |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
451 |
assume "indicator A x \<noteq> (0::pextreal)" |
39092 | 452 |
then have "x \<in> A" by (auto simp: indicator_def split: split_if_asm) |
453 |
moreover assume "x \<in> space M" "\<forall>y\<in>A. ?f x \<noteq> f y" |
|
454 |
ultimately show "f x = 0" by auto |
|
455 |
qed |
|
456 |
then show ?thesis by auto |
|
457 |
qed |
|
458 |
then show ?thesis |
|
459 |
unfolding simple_function_eq_borel_measurable |
|
460 |
R.simple_function_eq_borel_measurable |
|
461 |
unfolding borel_measurable_restricted[OF `A \<in> sets M`] |
|
462 |
by auto |
|
463 |
qed |
|
464 |
||
465 |
lemma (in sigma_algebra) simple_function_subalgebra: |
|
466 |
assumes "sigma_algebra.simple_function (M\<lparr>sets:=N\<rparr>) f" |
|
467 |
and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets:=N\<rparr>)" |
|
468 |
shows "simple_function f" |
|
469 |
using assms |
|
470 |
unfolding simple_function_def |
|
471 |
unfolding sigma_algebra.simple_function_def[OF N_subalgebra(2)] |
|
472 |
by auto |
|
473 |
||
40859 | 474 |
lemma (in sigma_algebra) simple_function_vimage: |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
475 |
fixes g :: "'a \<Rightarrow> pextreal" and f :: "'d \<Rightarrow> 'a" |
40859 | 476 |
assumes g: "simple_function g" and f: "f \<in> S \<rightarrow> space M" |
477 |
shows "sigma_algebra.simple_function (vimage_algebra S f) (\<lambda>x. g (f x))" |
|
478 |
proof - |
|
479 |
have subset: "(\<lambda>x. g (f x)) ` S \<subseteq> g ` space M" |
|
480 |
using f by auto |
|
481 |
interpret V: sigma_algebra "vimage_algebra S f" |
|
482 |
using f by (rule sigma_algebra_vimage) |
|
483 |
show ?thesis using g |
|
484 |
unfolding simple_function_eq_borel_measurable |
|
485 |
unfolding V.simple_function_eq_borel_measurable |
|
486 |
using measurable_vimage[OF _ f, of g borel] |
|
487 |
using finite_subset[OF subset] by auto |
|
488 |
qed |
|
489 |
||
38656 | 490 |
section "Simple integral" |
491 |
||
492 |
definition (in measure_space) |
|
493 |
"simple_integral f = (\<Sum>x \<in> f ` space M. x * \<mu> (f -` {x} \<inter> space M))" |
|
35582 | 494 |
|
38656 | 495 |
lemma (in measure_space) simple_integral_cong: |
496 |
assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t" |
|
497 |
shows "simple_integral f = simple_integral g" |
|
498 |
proof - |
|
499 |
have "f ` space M = g ` space M" |
|
500 |
"\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M" |
|
501 |
using assms by (auto intro!: image_eqI) |
|
502 |
thus ?thesis unfolding simple_integral_def by simp |
|
503 |
qed |
|
504 |
||
40859 | 505 |
lemma (in measure_space) simple_integral_cong_measure: |
506 |
assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A" and "simple_function f" |
|
507 |
shows "measure_space.simple_integral M \<nu> f = simple_integral f" |
|
508 |
proof - |
|
509 |
interpret v: measure_space M \<nu> |
|
510 |
by (rule measure_space_cong) fact |
|
40871 | 511 |
from simple_functionD[OF `simple_function f`] assms show ?thesis |
40859 | 512 |
unfolding simple_integral_def v.simple_integral_def |
513 |
by (auto intro!: setsum_cong) |
|
514 |
qed |
|
515 |
||
38656 | 516 |
lemma (in measure_space) simple_integral_const[simp]: |
517 |
"simple_integral (\<lambda>x. c) = c * \<mu> (space M)" |
|
518 |
proof (cases "space M = {}") |
|
519 |
case True thus ?thesis unfolding simple_integral_def by simp |
|
520 |
next |
|
521 |
case False hence "(\<lambda>x. c) ` space M = {c}" by auto |
|
522 |
thus ?thesis unfolding simple_integral_def by simp |
|
35582 | 523 |
qed |
524 |
||
38656 | 525 |
lemma (in measure_space) simple_function_partition: |
526 |
assumes "simple_function f" and "simple_function g" |
|
39910 | 527 |
shows "simple_integral f = (\<Sum>A\<in>(\<lambda>x. f -` {f x} \<inter> g -` {g x} \<inter> space M) ` space M. the_elem (f`A) * \<mu> A)" |
38656 | 528 |
(is "_ = setsum _ (?p ` space M)") |
529 |
proof- |
|
530 |
let "?sub x" = "?p ` (f -` {x} \<inter> space M)" |
|
531 |
let ?SIGMA = "Sigma (f`space M) ?sub" |
|
35582 | 532 |
|
38656 | 533 |
have [intro]: |
534 |
"finite (f ` space M)" |
|
535 |
"finite (g ` space M)" |
|
536 |
using assms unfolding simple_function_def by simp_all |
|
537 |
||
538 |
{ fix A |
|
539 |
have "?p ` (A \<inter> space M) \<subseteq> |
|
540 |
(\<lambda>(x,y). f -` {x} \<inter> g -` {y} \<inter> space M) ` (f`space M \<times> g`space M)" |
|
541 |
by auto |
|
542 |
hence "finite (?p ` (A \<inter> space M))" |
|
40786
0a54cfc9add3
gave more standard finite set rules simp and intro attribute
nipkow
parents:
39910
diff
changeset
|
543 |
by (rule finite_subset) auto } |
38656 | 544 |
note this[intro, simp] |
35582 | 545 |
|
38656 | 546 |
{ fix x assume "x \<in> space M" |
547 |
have "\<Union>(?sub (f x)) = (f -` {f x} \<inter> space M)" by auto |
|
548 |
moreover { |
|
549 |
fix x y |
|
550 |
have *: "f -` {f x} \<inter> g -` {g x} \<inter> space M |
|
551 |
= (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)" by auto |
|
552 |
assume "x \<in> space M" "y \<in> space M" |
|
553 |
hence "f -` {f x} \<inter> g -` {g x} \<inter> space M \<in> sets M" |
|
554 |
using assms unfolding simple_function_def * by auto } |
|
555 |
ultimately |
|
556 |
have "\<mu> (f -` {f x} \<inter> space M) = setsum (\<mu>) (?sub (f x))" |
|
557 |
by (subst measure_finitely_additive) auto } |
|
558 |
hence "simple_integral f = (\<Sum>(x,A)\<in>?SIGMA. x * \<mu> A)" |
|
559 |
unfolding simple_integral_def |
|
560 |
by (subst setsum_Sigma[symmetric], |
|
561 |
auto intro!: setsum_cong simp: setsum_right_distrib[symmetric]) |
|
39910 | 562 |
also have "\<dots> = (\<Sum>A\<in>?p ` space M. the_elem (f`A) * \<mu> A)" |
38656 | 563 |
proof - |
564 |
have [simp]: "\<And>x. x \<in> space M \<Longrightarrow> f ` ?p x = {f x}" by (auto intro!: imageI) |
|
39910 | 565 |
have "(\<lambda>A. (the_elem (f ` A), A)) ` ?p ` space M |
38656 | 566 |
= (\<lambda>x. (f x, ?p x)) ` space M" |
567 |
proof safe |
|
568 |
fix x assume "x \<in> space M" |
|
39910 | 569 |
thus "(f x, ?p x) \<in> (\<lambda>A. (the_elem (f`A), A)) ` ?p ` space M" |
38656 | 570 |
by (auto intro!: image_eqI[of _ _ "?p x"]) |
571 |
qed auto |
|
572 |
thus ?thesis |
|
39910 | 573 |
apply (auto intro!: setsum_reindex_cong[of "\<lambda>A. (the_elem (f`A), A)"] inj_onI) |
38656 | 574 |
apply (rule_tac x="xa" in image_eqI) |
575 |
by simp_all |
|
576 |
qed |
|
577 |
finally show ?thesis . |
|
35582 | 578 |
qed |
579 |
||
38656 | 580 |
lemma (in measure_space) simple_integral_add[simp]: |
581 |
assumes "simple_function f" and "simple_function g" |
|
582 |
shows "simple_integral (\<lambda>x. f x + g x) = simple_integral f + simple_integral g" |
|
35582 | 583 |
proof - |
38656 | 584 |
{ fix x let ?S = "g -` {g x} \<inter> f -` {f x} \<inter> space M" |
585 |
assume "x \<in> space M" |
|
586 |
hence "(\<lambda>a. f a + g a) ` ?S = {f x + g x}" "f ` ?S = {f x}" "g ` ?S = {g x}" |
|
587 |
"(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M = ?S" |
|
588 |
by auto } |
|
589 |
thus ?thesis |
|
590 |
unfolding |
|
591 |
simple_function_partition[OF simple_function_add[OF assms] simple_function_Pair[OF assms]] |
|
592 |
simple_function_partition[OF `simple_function f` `simple_function g`] |
|
593 |
simple_function_partition[OF `simple_function g` `simple_function f`] |
|
594 |
apply (subst (3) Int_commute) |
|
595 |
by (auto simp add: field_simps setsum_addf[symmetric] intro!: setsum_cong) |
|
35582 | 596 |
qed |
597 |
||
38656 | 598 |
lemma (in measure_space) simple_integral_setsum[simp]: |
599 |
assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function (f i)" |
|
600 |
shows "simple_integral (\<lambda>x. \<Sum>i\<in>P. f i x) = (\<Sum>i\<in>P. simple_integral (f i))" |
|
601 |
proof cases |
|
602 |
assume "finite P" |
|
603 |
from this assms show ?thesis |
|
604 |
by induct (auto simp: simple_function_setsum simple_integral_add) |
|
605 |
qed auto |
|
606 |
||
607 |
lemma (in measure_space) simple_integral_mult[simp]: |
|
608 |
assumes "simple_function f" |
|
609 |
shows "simple_integral (\<lambda>x. c * f x) = c * simple_integral f" |
|
610 |
proof - |
|
611 |
note mult = simple_function_mult[OF simple_function_const[of c] assms] |
|
612 |
{ fix x let ?S = "f -` {f x} \<inter> (\<lambda>x. c * f x) -` {c * f x} \<inter> space M" |
|
613 |
assume "x \<in> space M" |
|
614 |
hence "(\<lambda>x. c * f x) ` ?S = {c * f x}" "f ` ?S = {f x}" |
|
615 |
by auto } |
|
616 |
thus ?thesis |
|
617 |
unfolding simple_function_partition[OF mult assms] |
|
618 |
simple_function_partition[OF assms mult] |
|
619 |
by (auto simp: setsum_right_distrib field_simps intro!: setsum_cong) |
|
35582 | 620 |
qed |
621 |
||
40871 | 622 |
lemma (in sigma_algebra) simple_function_If: |
623 |
assumes sf: "simple_function f" "simple_function g" and A: "A \<in> sets M" |
|
624 |
shows "simple_function (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function ?IF") |
|
625 |
proof - |
|
626 |
def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M" |
|
627 |
show ?thesis unfolding simple_function_def |
|
628 |
proof safe |
|
629 |
have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto |
|
630 |
from finite_subset[OF this] assms |
|
631 |
show "finite (?IF ` space M)" unfolding simple_function_def by auto |
|
632 |
next |
|
633 |
fix x assume "x \<in> space M" |
|
634 |
then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A |
|
635 |
then ((F (f x) \<inter> A) \<union> (G (f x) - (G (f x) \<inter> A))) |
|
636 |
else ((F (g x) \<inter> A) \<union> (G (g x) - (G (g x) \<inter> A))))" |
|
637 |
using sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def) |
|
638 |
have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M" |
|
639 |
unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto |
|
640 |
show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto |
|
641 |
qed |
|
642 |
qed |
|
643 |
||
40859 | 644 |
lemma (in measure_space) simple_integral_mono_AE: |
645 |
assumes "simple_function f" and "simple_function g" |
|
646 |
and mono: "AE x. f x \<le> g x" |
|
647 |
shows "simple_integral f \<le> simple_integral g" |
|
648 |
proof - |
|
649 |
let "?S x" = "(g -` {g x} \<inter> space M) \<inter> (f -` {f x} \<inter> space M)" |
|
650 |
have *: "\<And>x. g -` {g x} \<inter> f -` {f x} \<inter> space M = ?S x" |
|
651 |
"\<And>x. f -` {f x} \<inter> g -` {g x} \<inter> space M = ?S x" by auto |
|
652 |
show ?thesis |
|
653 |
unfolding * |
|
654 |
simple_function_partition[OF `simple_function f` `simple_function g`] |
|
655 |
simple_function_partition[OF `simple_function g` `simple_function f`] |
|
656 |
proof (safe intro!: setsum_mono) |
|
657 |
fix x assume "x \<in> space M" |
|
658 |
then have *: "f ` ?S x = {f x}" "g ` ?S x = {g x}" by auto |
|
659 |
show "the_elem (f`?S x) * \<mu> (?S x) \<le> the_elem (g`?S x) * \<mu> (?S x)" |
|
660 |
proof (cases "f x \<le> g x") |
|
661 |
case True then show ?thesis using * by (auto intro!: mult_right_mono) |
|
662 |
next |
|
663 |
case False |
|
664 |
obtain N where N: "{x\<in>space M. \<not> f x \<le> g x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0" |
|
665 |
using mono by (auto elim!: AE_E) |
|
666 |
have "?S x \<subseteq> N" using N `x \<in> space M` False by auto |
|
40871 | 667 |
moreover have "?S x \<in> sets M" using assms |
668 |
by (rule_tac Int) (auto intro!: simple_functionD) |
|
40859 | 669 |
ultimately have "\<mu> (?S x) \<le> \<mu> N" |
670 |
using `N \<in> sets M` by (auto intro!: measure_mono) |
|
671 |
then show ?thesis using `\<mu> N = 0` by auto |
|
672 |
qed |
|
673 |
qed |
|
674 |
qed |
|
675 |
||
38656 | 676 |
lemma (in measure_space) simple_integral_mono: |
677 |
assumes "simple_function f" and "simple_function g" |
|
678 |
and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x" |
|
679 |
shows "simple_integral f \<le> simple_integral g" |
|
40859 | 680 |
proof (rule simple_integral_mono_AE[OF assms(1, 2)]) |
681 |
show "AE x. f x \<le> g x" |
|
682 |
using mono by (rule AE_cong) auto |
|
35582 | 683 |
qed |
684 |
||
40859 | 685 |
lemma (in measure_space) simple_integral_cong_AE: |
686 |
assumes "simple_function f" "simple_function g" and "AE x. f x = g x" |
|
687 |
shows "simple_integral f = simple_integral g" |
|
688 |
using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) |
|
689 |
||
690 |
lemma (in measure_space) simple_integral_cong': |
|
691 |
assumes sf: "simple_function f" "simple_function g" |
|
692 |
and mea: "\<mu> {x\<in>space M. f x \<noteq> g x} = 0" |
|
693 |
shows "simple_integral f = simple_integral g" |
|
694 |
proof (intro simple_integral_cong_AE sf AE_I) |
|
695 |
show "\<mu> {x\<in>space M. f x \<noteq> g x} = 0" by fact |
|
696 |
show "{x \<in> space M. f x \<noteq> g x} \<in> sets M" |
|
697 |
using sf[THEN borel_measurable_simple_function] by auto |
|
698 |
qed simp |
|
699 |
||
38656 | 700 |
lemma (in measure_space) simple_integral_indicator: |
701 |
assumes "A \<in> sets M" |
|
702 |
assumes "simple_function f" |
|
703 |
shows "simple_integral (\<lambda>x. f x * indicator A x) = |
|
704 |
(\<Sum>x \<in> f ` space M. x * \<mu> (f -` {x} \<inter> space M \<inter> A))" |
|
705 |
proof cases |
|
706 |
assume "A = space M" |
|
707 |
moreover hence "simple_integral (\<lambda>x. f x * indicator A x) = simple_integral f" |
|
708 |
by (auto intro!: simple_integral_cong) |
|
709 |
moreover have "\<And>X. X \<inter> space M \<inter> space M = X \<inter> space M" by auto |
|
710 |
ultimately show ?thesis by (simp add: simple_integral_def) |
|
711 |
next |
|
712 |
assume "A \<noteq> space M" |
|
713 |
then obtain x where x: "x \<in> space M" "x \<notin> A" using sets_into_space[OF assms(1)] by auto |
|
714 |
have I: "(\<lambda>x. f x * indicator A x) ` space M = f ` A \<union> {0}" (is "?I ` _ = _") |
|
35582 | 715 |
proof safe |
38656 | 716 |
fix y assume "?I y \<notin> f ` A" hence "y \<notin> A" by auto thus "?I y = 0" by auto |
717 |
next |
|
718 |
fix y assume "y \<in> A" thus "f y \<in> ?I ` space M" |
|
719 |
using sets_into_space[OF assms(1)] by (auto intro!: image_eqI[of _ _ y]) |
|
720 |
next |
|
721 |
show "0 \<in> ?I ` space M" using x by (auto intro!: image_eqI[of _ _ x]) |
|
35582 | 722 |
qed |
38656 | 723 |
have *: "simple_integral (\<lambda>x. f x * indicator A x) = |
724 |
(\<Sum>x \<in> f ` space M \<union> {0}. x * \<mu> (f -` {x} \<inter> space M \<inter> A))" |
|
725 |
unfolding simple_integral_def I |
|
726 |
proof (rule setsum_mono_zero_cong_left) |
|
727 |
show "finite (f ` space M \<union> {0})" |
|
728 |
using assms(2) unfolding simple_function_def by auto |
|
729 |
show "f ` A \<union> {0} \<subseteq> f`space M \<union> {0}" |
|
730 |
using sets_into_space[OF assms(1)] by auto |
|
40859 | 731 |
have "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}" |
732 |
by (auto simp: image_iff) |
|
38656 | 733 |
thus "\<forall>i\<in>f ` space M \<union> {0} - (f ` A \<union> {0}). |
734 |
i * \<mu> (f -` {i} \<inter> space M \<inter> A) = 0" by auto |
|
735 |
next |
|
736 |
fix x assume "x \<in> f`A \<union> {0}" |
|
737 |
hence "x \<noteq> 0 \<Longrightarrow> ?I -` {x} \<inter> space M = f -` {x} \<inter> space M \<inter> A" |
|
738 |
by (auto simp: indicator_def split: split_if_asm) |
|
739 |
thus "x * \<mu> (?I -` {x} \<inter> space M) = |
|
740 |
x * \<mu> (f -` {x} \<inter> space M \<inter> A)" by (cases "x = 0") simp_all |
|
741 |
qed |
|
742 |
show ?thesis unfolding * |
|
743 |
using assms(2) unfolding simple_function_def |
|
744 |
by (auto intro!: setsum_mono_zero_cong_right) |
|
745 |
qed |
|
35582 | 746 |
|
38656 | 747 |
lemma (in measure_space) simple_integral_indicator_only[simp]: |
748 |
assumes "A \<in> sets M" |
|
749 |
shows "simple_integral (indicator A) = \<mu> A" |
|
750 |
proof cases |
|
751 |
assume "space M = {}" hence "A = {}" using sets_into_space[OF assms] by auto |
|
752 |
thus ?thesis unfolding simple_integral_def using `space M = {}` by auto |
|
753 |
next |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
754 |
assume "space M \<noteq> {}" hence "(\<lambda>x. 1) ` space M = {1::pextreal}" by auto |
38656 | 755 |
thus ?thesis |
756 |
using simple_integral_indicator[OF assms simple_function_const[of 1]] |
|
757 |
using sets_into_space[OF assms] |
|
758 |
by (auto intro!: arg_cong[where f="\<mu>"]) |
|
759 |
qed |
|
35582 | 760 |
|
38656 | 761 |
lemma (in measure_space) simple_integral_null_set: |
762 |
assumes "simple_function u" "N \<in> null_sets" |
|
763 |
shows "simple_integral (\<lambda>x. u x * indicator N x) = 0" |
|
764 |
proof - |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
765 |
have "AE x. indicator N x = (0 :: pextreal)" |
40859 | 766 |
using `N \<in> null_sets` by (auto simp: indicator_def intro!: AE_I[of _ N]) |
767 |
then have "simple_integral (\<lambda>x. u x * indicator N x) = simple_integral (\<lambda>x. 0)" |
|
768 |
using assms by (intro simple_integral_cong_AE) (auto intro!: AE_disjI2) |
|
769 |
then show ?thesis by simp |
|
38656 | 770 |
qed |
35582 | 771 |
|
40859 | 772 |
lemma (in measure_space) simple_integral_cong_AE_mult_indicator: |
773 |
assumes sf: "simple_function f" and eq: "AE x. x \<in> S" and "S \<in> sets M" |
|
774 |
shows "simple_integral f = simple_integral (\<lambda>x. f x * indicator S x)" |
|
775 |
proof (rule simple_integral_cong_AE) |
|
776 |
show "simple_function f" by fact |
|
777 |
show "simple_function (\<lambda>x. f x * indicator S x)" |
|
778 |
using sf `S \<in> sets M` by auto |
|
779 |
from eq show "AE x. f x = f x * indicator S x" |
|
780 |
by (rule AE_mp) simp |
|
35582 | 781 |
qed |
782 |
||
39092 | 783 |
lemma (in measure_space) simple_integral_restricted: |
784 |
assumes "A \<in> sets M" |
|
785 |
assumes sf: "simple_function (\<lambda>x. f x * indicator A x)" |
|
786 |
shows "measure_space.simple_integral (restricted_space A) \<mu> f = simple_integral (\<lambda>x. f x * indicator A x)" |
|
787 |
(is "_ = simple_integral ?f") |
|
788 |
unfolding measure_space.simple_integral_def[OF restricted_measure_space[OF `A \<in> sets M`]] |
|
789 |
unfolding simple_integral_def |
|
790 |
proof (simp, safe intro!: setsum_mono_zero_cong_left) |
|
791 |
from sf show "finite (?f ` space M)" |
|
792 |
unfolding simple_function_def by auto |
|
793 |
next |
|
794 |
fix x assume "x \<in> A" |
|
795 |
then show "f x \<in> ?f ` space M" |
|
796 |
using sets_into_space `A \<in> sets M` by (auto intro!: image_eqI[of _ _ x]) |
|
797 |
next |
|
798 |
fix x assume "x \<in> space M" "?f x \<notin> f`A" |
|
799 |
then have "x \<notin> A" by (auto simp: image_iff) |
|
800 |
then show "?f x * \<mu> (?f -` {?f x} \<inter> space M) = 0" by simp |
|
801 |
next |
|
802 |
fix x assume "x \<in> A" |
|
803 |
then have "f x \<noteq> 0 \<Longrightarrow> |
|
804 |
f -` {f x} \<inter> A = ?f -` {f x} \<inter> space M" |
|
805 |
using `A \<in> sets M` sets_into_space |
|
806 |
by (auto simp: indicator_def split: split_if_asm) |
|
807 |
then show "f x * \<mu> (f -` {f x} \<inter> A) = |
|
808 |
f x * \<mu> (?f -` {f x} \<inter> space M)" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
809 |
unfolding pextreal_mult_cancel_left by auto |
39092 | 810 |
qed |
811 |
||
812 |
lemma (in measure_space) simple_integral_subalgebra[simp]: |
|
813 |
assumes "measure_space (M\<lparr>sets := N\<rparr>) \<mu>" |
|
814 |
shows "measure_space.simple_integral (M\<lparr>sets := N\<rparr>) \<mu> = simple_integral" |
|
815 |
unfolding simple_integral_def_raw |
|
816 |
unfolding measure_space.simple_integral_def_raw[OF assms] by simp |
|
817 |
||
40859 | 818 |
lemma (in measure_space) simple_integral_vimage: |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
819 |
fixes g :: "'a \<Rightarrow> pextreal" and f :: "'d \<Rightarrow> 'a" |
40859 | 820 |
assumes f: "bij_betw f S (space M)" |
821 |
shows "simple_integral g = |
|
822 |
measure_space.simple_integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>x. g (f x))" |
|
823 |
(is "_ = measure_space.simple_integral ?T ?\<mu> _") |
|
824 |
proof - |
|
825 |
from f interpret T: measure_space ?T ?\<mu> by (rule measure_space_isomorphic) |
|
826 |
have surj: "f`S = space M" |
|
827 |
using f unfolding bij_betw_def by simp |
|
828 |
have *: "(\<lambda>x. g (f x)) ` S = g ` f ` S" by auto |
|
829 |
have **: "f`S = space M" using f unfolding bij_betw_def by auto |
|
830 |
{ fix x assume "x \<in> space M" |
|
831 |
have "(f ` ((\<lambda>x. g (f x)) -` {g x} \<inter> S)) = |
|
832 |
(f ` (f -` (g -` {g x}) \<inter> S))" by auto |
|
833 |
also have "f -` (g -` {g x}) \<inter> S = f -` (g -` {g x} \<inter> space M) \<inter> S" |
|
834 |
using f unfolding bij_betw_def by auto |
|
835 |
also have "(f ` (f -` (g -` {g x} \<inter> space M) \<inter> S)) = g -` {g x} \<inter> space M" |
|
836 |
using ** by (intro image_vimage_inter_eq) auto |
|
837 |
finally have "(f ` ((\<lambda>x. g (f x)) -` {g x} \<inter> S)) = g -` {g x} \<inter> space M" by auto } |
|
838 |
then show ?thesis using assms |
|
839 |
unfolding simple_integral_def T.simple_integral_def bij_betw_def |
|
840 |
by (auto simp add: * intro!: setsum_cong) |
|
841 |
qed |
|
842 |
||
35692 | 843 |
section "Continuous posititve integration" |
844 |
||
38656 | 845 |
definition (in measure_space) |
40873 | 846 |
"positive_integral f = SUPR {g. simple_function g \<and> g \<le> f} simple_integral" |
40872 | 847 |
|
38656 | 848 |
lemma (in measure_space) positive_integral_alt: |
849 |
"positive_integral f = |
|
40873 | 850 |
(SUPR {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g`space M} simple_integral)" (is "_ = ?alt") |
40872 | 851 |
proof (rule antisym SUP_leI) |
40873 | 852 |
show "positive_integral f \<le> ?alt" unfolding positive_integral_def |
40872 | 853 |
proof (safe intro!: SUP_leI) |
854 |
fix g assume g: "simple_function g" "g \<le> f" |
|
855 |
let ?G = "g -` {\<omega>} \<inter> space M" |
|
856 |
show "simple_integral g \<le> |
|
857 |
SUPR {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g ` space M} simple_integral" |
|
858 |
(is "simple_integral g \<le> SUPR ?A simple_integral") |
|
859 |
proof cases |
|
860 |
let ?g = "\<lambda>x. indicator (space M - ?G) x * g x" |
|
861 |
have g': "simple_function ?g" |
|
862 |
using g by (auto intro: simple_functionD) |
|
863 |
moreover |
|
864 |
assume "\<mu> ?G = 0" |
|
865 |
then have "AE x. g x = ?g x" using g |
|
866 |
by (intro AE_I[where N="?G"]) |
|
867 |
(auto intro: simple_functionD simp: indicator_def) |
|
868 |
with g(1) g' have "simple_integral g = simple_integral ?g" |
|
869 |
by (rule simple_integral_cong_AE) |
|
870 |
moreover have "?g \<le> g" by (auto simp: le_fun_def indicator_def) |
|
871 |
from this `g \<le> f` have "?g \<le> f" by (rule order_trans) |
|
872 |
moreover have "\<omega> \<notin> ?g ` space M" |
|
873 |
by (auto simp: indicator_def split: split_if_asm) |
|
874 |
ultimately show ?thesis by (auto intro!: le_SUPI) |
|
875 |
next |
|
876 |
assume "\<mu> ?G \<noteq> 0" |
|
877 |
then have "?G \<noteq> {}" by auto |
|
878 |
then have "\<omega> \<in> g`space M" by force |
|
879 |
then have "space M \<noteq> {}" by auto |
|
880 |
have "SUPR ?A simple_integral = \<omega>" |
|
881 |
proof (intro SUP_\<omega>[THEN iffD2] allI impI) |
|
882 |
fix x assume "x < \<omega>" |
|
883 |
then guess n unfolding less_\<omega>_Ex_of_nat .. note n = this |
|
884 |
then have "0 < n" by (intro neq0_conv[THEN iffD1] notI) simp |
|
885 |
let ?g = "\<lambda>x. (of_nat n / (if \<mu> ?G = \<omega> then 1 else \<mu> ?G)) * indicator ?G x" |
|
886 |
show "\<exists>i\<in>?A. x < simple_integral i" |
|
887 |
proof (intro bexI impI CollectI conjI) |
|
888 |
show "simple_function ?g" using g |
|
889 |
by (auto intro!: simple_functionD simple_function_add) |
|
890 |
have "?g \<le> g" by (auto simp: le_fun_def indicator_def) |
|
891 |
from this g(2) show "?g \<le> f" by (rule order_trans) |
|
892 |
show "\<omega> \<notin> ?g ` space M" |
|
893 |
using `\<mu> ?G \<noteq> 0` by (auto simp: indicator_def split: split_if_asm) |
|
894 |
have "x < (of_nat n / (if \<mu> ?G = \<omega> then 1 else \<mu> ?G)) * \<mu> ?G" |
|
895 |
using n `\<mu> ?G \<noteq> 0` `0 < n` |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
896 |
by (auto simp: pextreal_noteq_omega_Ex field_simps) |
40872 | 897 |
also have "\<dots> = simple_integral ?g" using g `space M \<noteq> {}` |
898 |
by (subst simple_integral_indicator) |
|
899 |
(auto simp: image_constant ac_simps dest: simple_functionD) |
|
900 |
finally show "x < simple_integral ?g" . |
|
901 |
qed |
|
902 |
qed |
|
903 |
then show ?thesis by simp |
|
904 |
qed |
|
35582 | 905 |
qed |
40872 | 906 |
qed (auto intro!: SUP_subset simp: positive_integral_def) |
35582 | 907 |
|
40873 | 908 |
lemma (in measure_space) positive_integral_cong_measure: |
909 |
assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A" |
|
910 |
shows "measure_space.positive_integral M \<nu> f = positive_integral f" |
|
911 |
proof - |
|
912 |
interpret v: measure_space M \<nu> |
|
913 |
by (rule measure_space_cong) fact |
|
914 |
with assms show ?thesis |
|
915 |
unfolding positive_integral_def v.positive_integral_def SUPR_def |
|
916 |
by (auto intro!: arg_cong[where f=Sup] image_cong |
|
917 |
simp: simple_integral_cong_measure[of \<nu>]) |
|
918 |
qed |
|
919 |
||
920 |
lemma (in measure_space) positive_integral_alt1: |
|
921 |
"positive_integral f = |
|
922 |
(SUP g : {g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}. simple_integral g)" |
|
923 |
unfolding positive_integral_alt SUPR_def |
|
924 |
proof (safe intro!: arg_cong[where f=Sup]) |
|
925 |
fix g let ?g = "\<lambda>x. if x \<in> space M then g x else f x" |
|
926 |
assume "simple_function g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>" |
|
927 |
hence "?g \<le> f" "simple_function ?g" "simple_integral g = simple_integral ?g" |
|
928 |
"\<omega> \<notin> g`space M" |
|
929 |
unfolding le_fun_def by (auto cong: simple_function_cong simple_integral_cong) |
|
930 |
thus "simple_integral g \<in> simple_integral ` {g. simple_function g \<and> g \<le> f \<and> \<omega> \<notin> g`space M}" |
|
931 |
by auto |
|
932 |
next |
|
933 |
fix g assume "simple_function g" "g \<le> f" "\<omega> \<notin> g`space M" |
|
934 |
hence "simple_function g" "\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>" |
|
935 |
by (auto simp add: le_fun_def image_iff) |
|
936 |
thus "simple_integral g \<in> simple_integral ` {g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> f x \<and> g x \<noteq> \<omega>)}" |
|
937 |
by auto |
|
938 |
qed |
|
939 |
||
38656 | 940 |
lemma (in measure_space) positive_integral_cong: |
941 |
assumes "\<And>x. x \<in> space M \<Longrightarrow> f x = g x" |
|
942 |
shows "positive_integral f = positive_integral g" |
|
943 |
proof - |
|
944 |
have "\<And>h. (\<forall>x\<in>space M. h x \<le> f x \<and> h x \<noteq> \<omega>) = (\<forall>x\<in>space M. h x \<le> g x \<and> h x \<noteq> \<omega>)" |
|
945 |
using assms by auto |
|
946 |
thus ?thesis unfolding positive_integral_alt1 by auto |
|
947 |
qed |
|
948 |
||
949 |
lemma (in measure_space) positive_integral_eq_simple_integral: |
|
950 |
assumes "simple_function f" |
|
951 |
shows "positive_integral f = simple_integral f" |
|
40873 | 952 |
unfolding positive_integral_def |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
953 |
proof (safe intro!: pextreal_SUPI) |
38656 | 954 |
fix g assume "simple_function g" "g \<le> f" |
955 |
with assms show "simple_integral g \<le> simple_integral f" |
|
956 |
by (auto intro!: simple_integral_mono simp: le_fun_def) |
|
957 |
next |
|
958 |
fix y assume "\<forall>x. x\<in>{g. simple_function g \<and> g \<le> f} \<longrightarrow> simple_integral x \<le> y" |
|
959 |
with assms show "simple_integral f \<le> y" by auto |
|
960 |
qed |
|
35582 | 961 |
|
40859 | 962 |
lemma (in measure_space) positive_integral_mono_AE: |
963 |
assumes ae: "AE x. u x \<le> v x" |
|
38656 | 964 |
shows "positive_integral u \<le> positive_integral v" |
965 |
unfolding positive_integral_alt1 |
|
966 |
proof (safe intro!: SUPR_mono) |
|
40859 | 967 |
fix a assume a: "simple_function a" and mono: "\<forall>x\<in>space M. a x \<le> u x \<and> a x \<noteq> \<omega>" |
968 |
from ae obtain N where N: "{x\<in>space M. \<not> u x \<le> v x} \<subseteq> N" "N \<in> sets M" "\<mu> N = 0" |
|
969 |
by (auto elim!: AE_E) |
|
970 |
have "simple_function (\<lambda>x. a x * indicator (space M - N) x)" |
|
971 |
using `N \<in> sets M` a by auto |
|
972 |
with a show "\<exists>b\<in>{g. simple_function g \<and> (\<forall>x\<in>space M. g x \<le> v x \<and> g x \<noteq> \<omega>)}. |
|
973 |
simple_integral a \<le> simple_integral b" |
|
974 |
proof (safe intro!: bexI[of _ "\<lambda>x. a x * indicator (space M - N) x"] |
|
975 |
simple_integral_mono_AE) |
|
976 |
show "AE x. a x \<le> a x * indicator (space M - N) x" |
|
977 |
proof (rule AE_I, rule subset_refl) |
|
978 |
have *: "{x \<in> space M. \<not> a x \<le> a x * indicator (space M - N) x} = |
|
979 |
N \<inter> {x \<in> space M. a x \<noteq> 0}" (is "?N = _") |
|
980 |
using `N \<in> sets M`[THEN sets_into_space] by (auto simp: indicator_def) |
|
981 |
then show "?N \<in> sets M" |
|
982 |
using `N \<in> sets M` `simple_function a`[THEN borel_measurable_simple_function] |
|
983 |
by (auto intro!: measure_mono Int) |
|
984 |
then have "\<mu> ?N \<le> \<mu> N" |
|
985 |
unfolding * using `N \<in> sets M` by (auto intro!: measure_mono) |
|
986 |
then show "\<mu> ?N = 0" using `\<mu> N = 0` by auto |
|
987 |
qed |
|
988 |
next |
|
989 |
fix x assume "x \<in> space M" |
|
990 |
show "a x * indicator (space M - N) x \<le> v x" |
|
991 |
proof (cases "x \<in> N") |
|
992 |
case True then show ?thesis by simp |
|
993 |
next |
|
994 |
case False |
|
995 |
with N mono have "a x \<le> u x" "u x \<le> v x" using `x \<in> space M` by auto |
|
996 |
with False `x \<in> space M` show "a x * indicator (space M - N) x \<le> v x" by auto |
|
997 |
qed |
|
998 |
assume "a x * indicator (space M - N) x = \<omega>" |
|
999 |
with mono `x \<in> space M` show False |
|
1000 |
by (simp split: split_if_asm add: indicator_def) |
|
1001 |
qed |
|
1002 |
qed |
|
1003 |
||
1004 |
lemma (in measure_space) positive_integral_cong_AE: |
|
1005 |
"AE x. u x = v x \<Longrightarrow> positive_integral u = positive_integral v" |
|
1006 |
by (auto simp: eq_iff intro!: positive_integral_mono_AE) |
|
1007 |
||
1008 |
lemma (in measure_space) positive_integral_mono: |
|
1009 |
assumes mono: "\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x" |
|
1010 |
shows "positive_integral u \<le> positive_integral v" |
|
1011 |
using mono by (auto intro!: AE_cong positive_integral_mono_AE) |
|
1012 |
||
40873 | 1013 |
lemma image_set_cong: |
1014 |
assumes A: "\<And>x. x \<in> A \<Longrightarrow> \<exists>y\<in>B. f x = g y" |
|
1015 |
assumes B: "\<And>y. y \<in> B \<Longrightarrow> \<exists>x\<in>A. g y = f x" |
|
1016 |
shows "f ` A = g ` B" |
|
1017 |
using assms by blast |
|
1018 |
||
40859 | 1019 |
lemma (in measure_space) positive_integral_vimage: |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1020 |
fixes g :: "'a \<Rightarrow> pextreal" and f :: "'d \<Rightarrow> 'a" |
40859 | 1021 |
assumes f: "bij_betw f S (space M)" |
1022 |
shows "positive_integral g = |
|
1023 |
measure_space.positive_integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>x. g (f x))" |
|
1024 |
(is "_ = measure_space.positive_integral ?T ?\<mu> _") |
|
1025 |
proof - |
|
1026 |
from f interpret T: measure_space ?T ?\<mu> by (rule measure_space_isomorphic) |
|
1027 |
have f_fun: "f \<in> S \<rightarrow> space M" using assms unfolding bij_betw_def by auto |
|
1028 |
from assms have inv: "bij_betw (the_inv_into S f) (space M) S" |
|
1029 |
by (rule bij_betw_the_inv_into) |
|
1030 |
then have inv_fun: "the_inv_into S f \<in> space M \<rightarrow> S" unfolding bij_betw_def by auto |
|
1031 |
have surj: "f`S = space M" |
|
1032 |
using f unfolding bij_betw_def by simp |
|
1033 |
have inj: "inj_on f S" |
|
1034 |
using f unfolding bij_betw_def by simp |
|
1035 |
have inv_f: "\<And>x. x \<in> space M \<Longrightarrow> f (the_inv_into S f x) = x" |
|
1036 |
using f_the_inv_into_f[of f S] f unfolding bij_betw_def by auto |
|
1037 |
from simple_integral_vimage[OF assms, symmetric] |
|
1038 |
have *: "simple_integral = T.simple_integral \<circ> (\<lambda>g. g \<circ> f)" by (simp add: comp_def) |
|
1039 |
show ?thesis |
|
1040 |
unfolding positive_integral_alt1 T.positive_integral_alt1 SUPR_def * image_compose |
|
1041 |
proof (safe intro!: arg_cong[where f=Sup] image_set_cong, simp_all add: comp_def) |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1042 |
fix g' :: "'a \<Rightarrow> pextreal" assume "simple_function g'" "\<forall>x\<in>space M. g' x \<le> g x \<and> g' x \<noteq> \<omega>" |
40859 | 1043 |
then show "\<exists>h. T.simple_function h \<and> (\<forall>x\<in>S. h x \<le> g (f x) \<and> h x \<noteq> \<omega>) \<and> |
1044 |
T.simple_integral (\<lambda>x. g' (f x)) = T.simple_integral h" |
|
1045 |
using f unfolding bij_betw_def |
|
1046 |
by (auto intro!: exI[of _ "\<lambda>x. g' (f x)"] |
|
1047 |
simp add: le_fun_def simple_function_vimage[OF _ f_fun]) |
|
1048 |
next |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1049 |
fix g' :: "'d \<Rightarrow> pextreal" assume g': "T.simple_function g'" "\<forall>x\<in>S. g' x \<le> g (f x) \<and> g' x \<noteq> \<omega>" |
40859 | 1050 |
let ?g = "\<lambda>x. g' (the_inv_into S f x)" |
1051 |
show "\<exists>h. simple_function h \<and> (\<forall>x\<in>space M. h x \<le> g x \<and> h x \<noteq> \<omega>) \<and> |
|
1052 |
T.simple_integral g' = T.simple_integral (\<lambda>x. h (f x))" |
|
1053 |
proof (intro exI[of _ ?g] conjI ballI) |
|
1054 |
{ fix x assume x: "x \<in> space M" |
|
1055 |
then have "the_inv_into S f x \<in> S" using inv_fun by auto |
|
1056 |
with g' have "g' (the_inv_into S f x) \<le> g (f (the_inv_into S f x)) \<and> g' (the_inv_into S f x) \<noteq> \<omega>" |
|
1057 |
by auto |
|
1058 |
then show "g' (the_inv_into S f x) \<le> g x" "g' (the_inv_into S f x) \<noteq> \<omega>" |
|
1059 |
using f_the_inv_into_f[of f S x] x f unfolding bij_betw_def by auto } |
|
1060 |
note vimage_vimage_inv[OF f inv_f inv_fun, simp] |
|
1061 |
from T.simple_function_vimage[OF g'(1), unfolded space_vimage_algebra, OF inv_fun] |
|
1062 |
show "simple_function (\<lambda>x. g' (the_inv_into S f x))" |
|
1063 |
unfolding simple_function_def by (simp add: simple_function_def) |
|
1064 |
show "T.simple_integral g' = T.simple_integral (\<lambda>x. ?g (f x))" |
|
1065 |
using the_inv_into_f_f[OF inj] by (auto intro!: T.simple_integral_cong) |
|
1066 |
qed |
|
1067 |
qed |
|
1068 |
qed |
|
1069 |
||
1070 |
lemma (in measure_space) positive_integral_vimage_inv: |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1071 |
fixes g :: "'d \<Rightarrow> pextreal" and f :: "'d \<Rightarrow> 'a" |
40859 | 1072 |
assumes f: "bij_betw f S (space M)" |
1073 |
shows "measure_space.positive_integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) g = |
|
1074 |
positive_integral (\<lambda>x. g (the_inv_into S f x))" |
|
1075 |
proof - |
|
1076 |
interpret v: measure_space "vimage_algebra S f" "\<lambda>A. \<mu> (f ` A)" |
|
1077 |
using f by (rule measure_space_isomorphic) |
|
1078 |
show ?thesis |
|
1079 |
unfolding positive_integral_vimage[OF f, of "\<lambda>x. g (the_inv_into S f x)"] |
|
1080 |
using f[unfolded bij_betw_def] |
|
1081 |
by (auto intro!: v.positive_integral_cong simp: the_inv_into_f_f) |
|
38656 | 1082 |
qed |
1083 |
||
1084 |
lemma (in measure_space) positive_integral_SUP_approx: |
|
1085 |
assumes "f \<up> s" |
|
1086 |
and f: "\<And>i. f i \<in> borel_measurable M" |
|
1087 |
and "simple_function u" |
|
1088 |
and le: "u \<le> s" and real: "\<omega> \<notin> u`space M" |
|
1089 |
shows "simple_integral u \<le> (SUP i. positive_integral (f i))" (is "_ \<le> ?S") |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1090 |
proof (rule pextreal_le_mult_one_interval) |
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1091 |
fix a :: pextreal assume "0 < a" "a < 1" |
38656 | 1092 |
hence "a \<noteq> 0" by auto |
1093 |
let "?B i" = "{x \<in> space M. a * u x \<le> f i x}" |
|
1094 |
have B: "\<And>i. ?B i \<in> sets M" |
|
1095 |
using f `simple_function u` by (auto simp: borel_measurable_simple_function) |
|
1096 |
||
1097 |
let "?uB i x" = "u x * indicator (?B i) x" |
|
1098 |
||
1099 |
{ fix i have "?B i \<subseteq> ?B (Suc i)" |
|
1100 |
proof safe |
|
1101 |
fix i x assume "a * u x \<le> f i x" |
|
1102 |
also have "\<dots> \<le> f (Suc i) x" |
|
1103 |
using `f \<up> s` unfolding isoton_def le_fun_def by auto |
|
1104 |
finally show "a * u x \<le> f (Suc i) x" . |
|
1105 |
qed } |
|
1106 |
note B_mono = this |
|
35582 | 1107 |
|
38656 | 1108 |
have u: "\<And>x. x \<in> space M \<Longrightarrow> u -` {u x} \<inter> space M \<in> sets M" |
1109 |
using `simple_function u` by (auto simp add: simple_function_def) |
|
1110 |
||
40859 | 1111 |
have "\<And>i. (\<lambda>n. (u -` {i} \<inter> space M) \<inter> ?B n) \<up> (u -` {i} \<inter> space M)" using B_mono unfolding isoton_def |
1112 |
proof safe |
|
1113 |
fix x i assume "x \<in> space M" |
|
1114 |
show "x \<in> (\<Union>i. (u -` {u x} \<inter> space M) \<inter> ?B i)" |
|
1115 |
proof cases |
|
1116 |
assume "u x = 0" thus ?thesis using `x \<in> space M` by simp |
|
1117 |
next |
|
1118 |
assume "u x \<noteq> 0" |
|
1119 |
with `a < 1` real `x \<in> space M` |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1120 |
have "a * u x < 1 * u x" by (rule_tac pextreal_mult_strict_right_mono) (auto simp: image_iff) |
40859 | 1121 |
also have "\<dots> \<le> (SUP i. f i x)" using le `f \<up> s` |
1122 |
unfolding isoton_fun_expand by (auto simp: isoton_def le_fun_def) |
|
1123 |
finally obtain i where "a * u x < f i x" unfolding SUPR_def |
|
1124 |
by (auto simp add: less_Sup_iff) |
|
1125 |
hence "a * u x \<le> f i x" by auto |
|
1126 |
thus ?thesis using `x \<in> space M` by auto |
|
1127 |
qed |
|
1128 |
qed auto |
|
1129 |
note measure_conv = measure_up[OF Int[OF u B] this] |
|
38656 | 1130 |
|
1131 |
have "simple_integral u = (SUP i. simple_integral (?uB i))" |
|
1132 |
unfolding simple_integral_indicator[OF B `simple_function u`] |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1133 |
proof (subst SUPR_pextreal_setsum, safe) |
38656 | 1134 |
fix x n assume "x \<in> space M" |
1135 |
have "\<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f n x}) |
|
1136 |
\<le> \<mu> (u -` {u x} \<inter> space M \<inter> {x \<in> space M. a * u x \<le> f (Suc n) x})" |
|
1137 |
using B_mono Int[OF u[OF `x \<in> space M`] B] by (auto intro!: measure_mono) |
|
1138 |
thus "u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B n) |
|
1139 |
\<le> u x * \<mu> (u -` {u x} \<inter> space M \<inter> ?B (Suc n))" |
|
1140 |
by (auto intro: mult_left_mono) |
|
1141 |
next |
|
1142 |
show "simple_integral u = |
|
1143 |
(\<Sum>i\<in>u ` space M. SUP n. i * \<mu> (u -` {i} \<inter> space M \<inter> ?B n))" |
|
1144 |
using measure_conv unfolding simple_integral_def isoton_def |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1145 |
by (auto intro!: setsum_cong simp: pextreal_SUP_cmult) |
38656 | 1146 |
qed |
1147 |
moreover |
|
1148 |
have "a * (SUP i. simple_integral (?uB i)) \<le> ?S" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1149 |
unfolding pextreal_SUP_cmult[symmetric] |
38705 | 1150 |
proof (safe intro!: SUP_mono bexI) |
38656 | 1151 |
fix i |
1152 |
have "a * simple_integral (?uB i) = simple_integral (\<lambda>x. a * ?uB i x)" |
|
1153 |
using B `simple_function u` |
|
1154 |
by (subst simple_integral_mult[symmetric]) (auto simp: field_simps) |
|
1155 |
also have "\<dots> \<le> positive_integral (f i)" |
|
1156 |
proof - |
|
1157 |
have "\<And>x. a * ?uB i x \<le> f i x" unfolding indicator_def by auto |
|
1158 |
hence *: "simple_function (\<lambda>x. a * ?uB i x)" using B assms(3) |
|
1159 |
by (auto intro!: simple_integral_mono) |
|
1160 |
show ?thesis unfolding positive_integral_eq_simple_integral[OF *, symmetric] |
|
1161 |
by (auto intro!: positive_integral_mono simp: indicator_def) |
|
1162 |
qed |
|
1163 |
finally show "a * simple_integral (?uB i) \<le> positive_integral (f i)" |
|
1164 |
by auto |
|
38705 | 1165 |
qed simp |
38656 | 1166 |
ultimately show "a * simple_integral u \<le> ?S" by simp |
35582 | 1167 |
qed |
1168 |
||
1169 |
text {* Beppo-Levi monotone convergence theorem *} |
|
38656 | 1170 |
lemma (in measure_space) positive_integral_isoton: |
1171 |
assumes "f \<up> u" "\<And>i. f i \<in> borel_measurable M" |
|
1172 |
shows "(\<lambda>i. positive_integral (f i)) \<up> positive_integral u" |
|
1173 |
unfolding isoton_def |
|
1174 |
proof safe |
|
1175 |
fix i show "positive_integral (f i) \<le> positive_integral (f (Suc i))" |
|
1176 |
apply (rule positive_integral_mono) |
|
1177 |
using `f \<up> u` unfolding isoton_def le_fun_def by auto |
|
1178 |
next |
|
1179 |
have "u \<in> borel_measurable M" |
|
1180 |
using borel_measurable_SUP[of UNIV f] assms by (auto simp: isoton_def) |
|
1181 |
have u: "u = (SUP i. f i)" using `f \<up> u` unfolding isoton_def by auto |
|
35582 | 1182 |
|
38656 | 1183 |
show "(SUP i. positive_integral (f i)) = positive_integral u" |
1184 |
proof (rule antisym) |
|
1185 |
from `f \<up> u`[THEN isoton_Sup, unfolded le_fun_def] |
|
1186 |
show "(SUP j. positive_integral (f j)) \<le> positive_integral u" |
|
1187 |
by (auto intro!: SUP_leI positive_integral_mono) |
|
1188 |
next |
|
1189 |
show "positive_integral u \<le> (SUP i. positive_integral (f i))" |
|
40873 | 1190 |
unfolding positive_integral_alt[of u] |
38656 | 1191 |
by (auto intro!: SUP_leI positive_integral_SUP_approx[OF assms]) |
35582 | 1192 |
qed |
1193 |
qed |
|
1194 |
||
40859 | 1195 |
lemma (in measure_space) positive_integral_monotone_convergence_SUP: |
1196 |
assumes "\<And>i x. x \<in> space M \<Longrightarrow> f i x \<le> f (Suc i) x" |
|
1197 |
assumes "\<And>i. f i \<in> borel_measurable M" |
|
1198 |
shows "(SUP i. positive_integral (f i)) = positive_integral (\<lambda>x. SUP i. f i x)" |
|
1199 |
(is "_ = positive_integral ?u") |
|
1200 |
proof - |
|
1201 |
have "?u \<in> borel_measurable M" |
|
1202 |
using borel_measurable_SUP[of _ f] assms by (simp add: SUPR_fun_expand) |
|
1203 |
show ?thesis |
|
1204 |
proof (rule antisym) |
|
1205 |
show "(SUP j. positive_integral (f j)) \<le> positive_integral ?u" |
|
1206 |
by (auto intro!: SUP_leI positive_integral_mono le_SUPI) |
|
1207 |
next |
|
1208 |
def rf \<equiv> "\<lambda>i. \<lambda>x\<in>space M. f i x" and ru \<equiv> "\<lambda>x\<in>space M. ?u x" |
|
1209 |
have "\<And>i. rf i \<in> borel_measurable M" unfolding rf_def |
|
1210 |
using assms by (simp cong: measurable_cong) |
|
1211 |
moreover have iso: "rf \<up> ru" using assms unfolding rf_def ru_def |
|
1212 |
unfolding isoton_def SUPR_fun_expand le_fun_def fun_eq_iff |
|
40872 | 1213 |
using SUP_const[OF UNIV_not_empty] |
40859 | 1214 |
by (auto simp: restrict_def le_fun_def SUPR_fun_expand fun_eq_iff) |
1215 |
ultimately have "positive_integral ru \<le> (SUP i. positive_integral (rf i))" |
|
40873 | 1216 |
unfolding positive_integral_alt[of ru] |
40859 | 1217 |
by (auto simp: le_fun_def intro!: SUP_leI positive_integral_SUP_approx) |
1218 |
then show "positive_integral ?u \<le> (SUP i. positive_integral (f i))" |
|
1219 |
unfolding ru_def rf_def by (simp cong: positive_integral_cong) |
|
1220 |
qed |
|
1221 |
qed |
|
1222 |
||
38656 | 1223 |
lemma (in measure_space) SUP_simple_integral_sequences: |
1224 |
assumes f: "f \<up> u" "\<And>i. simple_function (f i)" |
|
1225 |
and g: "g \<up> u" "\<And>i. simple_function (g i)" |
|
1226 |
shows "(SUP i. simple_integral (f i)) = (SUP i. simple_integral (g i))" |
|
1227 |
(is "SUPR _ ?F = SUPR _ ?G") |
|
1228 |
proof - |
|
1229 |
have "(SUP i. ?F i) = (SUP i. positive_integral (f i))" |
|
1230 |
using assms by (simp add: positive_integral_eq_simple_integral) |
|
1231 |
also have "\<dots> = positive_integral u" |
|
1232 |
using positive_integral_isoton[OF f(1) borel_measurable_simple_function[OF f(2)]] |
|
1233 |
unfolding isoton_def by simp |
|
1234 |
also have "\<dots> = (SUP i. positive_integral (g i))" |
|
1235 |
using positive_integral_isoton[OF g(1) borel_measurable_simple_function[OF g(2)]] |
|
1236 |
unfolding isoton_def by simp |
|
1237 |
also have "\<dots> = (SUP i. ?G i)" |
|
1238 |
using assms by (simp add: positive_integral_eq_simple_integral) |
|
1239 |
finally show ?thesis . |
|
1240 |
qed |
|
1241 |
||
1242 |
lemma (in measure_space) positive_integral_const[simp]: |
|
1243 |
"positive_integral (\<lambda>x. c) = c * \<mu> (space M)" |
|
1244 |
by (subst positive_integral_eq_simple_integral) auto |
|
1245 |
||
1246 |
lemma (in measure_space) positive_integral_isoton_simple: |
|
1247 |
assumes "f \<up> u" and e: "\<And>i. simple_function (f i)" |
|
1248 |
shows "(\<lambda>i. simple_integral (f i)) \<up> positive_integral u" |
|
1249 |
using positive_integral_isoton[OF `f \<up> u` e(1)[THEN borel_measurable_simple_function]] |
|
1250 |
unfolding positive_integral_eq_simple_integral[OF e] . |
|
1251 |
||
1252 |
lemma (in measure_space) positive_integral_linear: |
|
1253 |
assumes f: "f \<in> borel_measurable M" |
|
1254 |
and g: "g \<in> borel_measurable M" |
|
1255 |
shows "positive_integral (\<lambda>x. a * f x + g x) = |
|
1256 |
a * positive_integral f + positive_integral g" |
|
1257 |
(is "positive_integral ?L = _") |
|
35582 | 1258 |
proof - |
38656 | 1259 |
from borel_measurable_implies_simple_function_sequence'[OF f] guess u . |
1260 |
note u = this positive_integral_isoton_simple[OF this(1-2)] |
|
1261 |
from borel_measurable_implies_simple_function_sequence'[OF g] guess v . |
|
1262 |
note v = this positive_integral_isoton_simple[OF this(1-2)] |
|
1263 |
let "?L' i x" = "a * u i x + v i x" |
|
1264 |
||
1265 |
have "?L \<in> borel_measurable M" |
|
1266 |
using assms by simp |
|
1267 |
from borel_measurable_implies_simple_function_sequence'[OF this] guess l . |
|
1268 |
note positive_integral_isoton_simple[OF this(1-2)] and l = this |
|
1269 |
moreover have |
|
1270 |
"(SUP i. simple_integral (l i)) = (SUP i. simple_integral (?L' i))" |
|
1271 |
proof (rule SUP_simple_integral_sequences[OF l(1-2)]) |
|
1272 |
show "?L' \<up> ?L" "\<And>i. simple_function (?L' i)" |
|
1273 |
using u v by (auto simp: isoton_fun_expand isoton_add isoton_cmult_right) |
|
1274 |
qed |
|
1275 |
moreover from u v have L'_isoton: |
|
1276 |
"(\<lambda>i. simple_integral (?L' i)) \<up> a * positive_integral f + positive_integral g" |
|
1277 |
by (simp add: isoton_add isoton_cmult_right) |
|
1278 |
ultimately show ?thesis by (simp add: isoton_def) |
|
1279 |
qed |
|
1280 |
||
1281 |
lemma (in measure_space) positive_integral_cmult: |
|
1282 |
assumes "f \<in> borel_measurable M" |
|
1283 |
shows "positive_integral (\<lambda>x. c * f x) = c * positive_integral f" |
|
1284 |
using positive_integral_linear[OF assms, of "\<lambda>x. 0" c] by auto |
|
1285 |
||
1286 |
lemma (in measure_space) positive_integral_indicator[simp]: |
|
1287 |
"A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. indicator A x) = \<mu> A" |
|
1288 |
by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator) |
|
1289 |
||
1290 |
lemma (in measure_space) positive_integral_cmult_indicator: |
|
1291 |
"A \<in> sets M \<Longrightarrow> positive_integral (\<lambda>x. c * indicator A x) = c * \<mu> A" |
|
1292 |
by (subst positive_integral_eq_simple_integral) (auto simp: simple_function_indicator simple_integral_indicator) |
|
1293 |
||
1294 |
lemma (in measure_space) positive_integral_add: |
|
1295 |
assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
|
1296 |
shows "positive_integral (\<lambda>x. f x + g x) = positive_integral f + positive_integral g" |
|
1297 |
using positive_integral_linear[OF assms, of 1] by simp |
|
1298 |
||
1299 |
lemma (in measure_space) positive_integral_setsum: |
|
1300 |
assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M" |
|
1301 |
shows "positive_integral (\<lambda>x. \<Sum>i\<in>P. f i x) = (\<Sum>i\<in>P. positive_integral (f i))" |
|
1302 |
proof cases |
|
1303 |
assume "finite P" |
|
1304 |
from this assms show ?thesis |
|
1305 |
proof induct |
|
1306 |
case (insert i P) |
|
1307 |
have "f i \<in> borel_measurable M" |
|
1308 |
"(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1309 |
using insert by (auto intro!: borel_measurable_pextreal_setsum) |
38656 | 1310 |
from positive_integral_add[OF this] |
1311 |
show ?case using insert by auto |
|
1312 |
qed simp |
|
1313 |
qed simp |
|
1314 |
||
1315 |
lemma (in measure_space) positive_integral_diff: |
|
1316 |
assumes f: "f \<in> borel_measurable M" and g: "g \<in> borel_measurable M" |
|
1317 |
and fin: "positive_integral g \<noteq> \<omega>" |
|
1318 |
and mono: "\<And>x. x \<in> space M \<Longrightarrow> g x \<le> f x" |
|
1319 |
shows "positive_integral (\<lambda>x. f x - g x) = positive_integral f - positive_integral g" |
|
1320 |
proof - |
|
1321 |
have borel: "(\<lambda>x. f x - g x) \<in> borel_measurable M" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1322 |
using f g by (rule borel_measurable_pextreal_diff) |
38656 | 1323 |
have "positive_integral (\<lambda>x. f x - g x) + positive_integral g = |
1324 |
positive_integral f" |
|
1325 |
unfolding positive_integral_add[OF borel g, symmetric] |
|
1326 |
proof (rule positive_integral_cong) |
|
1327 |
fix x assume "x \<in> space M" |
|
1328 |
from mono[OF this] show "f x - g x + g x = f x" |
|
1329 |
by (cases "f x", cases "g x", simp, simp, cases "g x", auto) |
|
1330 |
qed |
|
1331 |
with mono show ?thesis |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1332 |
by (subst minus_pextreal_eq2[OF _ fin]) (auto intro!: positive_integral_mono) |
38656 | 1333 |
qed |
1334 |
||
1335 |
lemma (in measure_space) positive_integral_psuminf: |
|
1336 |
assumes "\<And>i. f i \<in> borel_measurable M" |
|
1337 |
shows "positive_integral (\<lambda>x. \<Sum>\<^isub>\<infinity> i. f i x) = (\<Sum>\<^isub>\<infinity> i. positive_integral (f i))" |
|
1338 |
proof - |
|
1339 |
have "(\<lambda>i. positive_integral (\<lambda>x. \<Sum>i<i. f i x)) \<up> positive_integral (\<lambda>x. \<Sum>\<^isub>\<infinity>i. f i x)" |
|
1340 |
by (rule positive_integral_isoton) |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1341 |
(auto intro!: borel_measurable_pextreal_setsum assms positive_integral_mono |
38656 | 1342 |
arg_cong[where f=Sup] |
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
1343 |
simp: isoton_def le_fun_def psuminf_def fun_eq_iff SUPR_def Sup_fun_def) |
38656 | 1344 |
thus ?thesis |
1345 |
by (auto simp: isoton_def psuminf_def positive_integral_setsum[OF assms]) |
|
1346 |
qed |
|
1347 |
||
1348 |
text {* Fatou's lemma: convergence theorem on limes inferior *} |
|
1349 |
lemma (in measure_space) positive_integral_lim_INF: |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1350 |
fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> pextreal" |
38656 | 1351 |
assumes "\<And>i. u i \<in> borel_measurable M" |
1352 |
shows "positive_integral (SUP n. INF m. u (m + n)) \<le> |
|
1353 |
(SUP n. INF m. positive_integral (u (m + n)))" |
|
1354 |
proof - |
|
1355 |
have "(SUP n. INF m. u (m + n)) \<in> borel_measurable M" |
|
1356 |
by (auto intro!: borel_measurable_SUP borel_measurable_INF assms) |
|
1357 |
||
1358 |
have "(\<lambda>n. INF m. u (m + n)) \<up> (SUP n. INF m. u (m + n))" |
|
38705 | 1359 |
proof (unfold isoton_def, safe intro!: INF_mono bexI) |
1360 |
fix i m show "u (Suc m + i) \<le> u (m + Suc i)" by simp |
|
1361 |
qed simp |
|
38656 | 1362 |
from positive_integral_isoton[OF this] assms |
1363 |
have "positive_integral (SUP n. INF m. u (m + n)) = |
|
1364 |
(SUP n. positive_integral (INF m. u (m + n)))" |
|
1365 |
unfolding isoton_def by (simp add: borel_measurable_INF) |
|
1366 |
also have "\<dots> \<le> (SUP n. INF m. positive_integral (u (m + n)))" |
|
38705 | 1367 |
apply (rule SUP_mono) |
1368 |
apply (rule_tac x=n in bexI) |
|
1369 |
by (auto intro!: positive_integral_mono INFI_bound INF_leI exI simp: INFI_fun_expand) |
|
38656 | 1370 |
finally show ?thesis . |
35582 | 1371 |
qed |
1372 |
||
38656 | 1373 |
lemma (in measure_space) measure_space_density: |
1374 |
assumes borel: "u \<in> borel_measurable M" |
|
1375 |
shows "measure_space M (\<lambda>A. positive_integral (\<lambda>x. u x * indicator A x))" (is "measure_space M ?v") |
|
1376 |
proof |
|
1377 |
show "?v {} = 0" by simp |
|
1378 |
show "countably_additive M ?v" |
|
1379 |
unfolding countably_additive_def |
|
1380 |
proof safe |
|
1381 |
fix A :: "nat \<Rightarrow> 'a set" |
|
1382 |
assume "range A \<subseteq> sets M" |
|
1383 |
hence "\<And>i. (\<lambda>x. u x * indicator (A i) x) \<in> borel_measurable M" |
|
1384 |
using borel by (auto intro: borel_measurable_indicator) |
|
1385 |
moreover assume "disjoint_family A" |
|
1386 |
note psuminf_indicator[OF this] |
|
1387 |
ultimately show "(\<Sum>\<^isub>\<infinity>n. ?v (A n)) = ?v (\<Union>x. A x)" |
|
1388 |
by (simp add: positive_integral_psuminf[symmetric]) |
|
1389 |
qed |
|
1390 |
qed |
|
35582 | 1391 |
|
39092 | 1392 |
lemma (in measure_space) positive_integral_translated_density: |
1393 |
assumes "f \<in> borel_measurable M" "g \<in> borel_measurable M" |
|
1394 |
shows "measure_space.positive_integral M (\<lambda>A. positive_integral (\<lambda>x. f x * indicator A x)) g = |
|
1395 |
positive_integral (\<lambda>x. f x * g x)" (is "measure_space.positive_integral M ?T _ = _") |
|
1396 |
proof - |
|
1397 |
from measure_space_density[OF assms(1)] |
|
1398 |
interpret T: measure_space M ?T . |
|
1399 |
from borel_measurable_implies_simple_function_sequence[OF assms(2)] |
|
1400 |
obtain G where G: "\<And>i. simple_function (G i)" "G \<up> g" by blast |
|
1401 |
note G_borel = borel_measurable_simple_function[OF this(1)] |
|
1402 |
from T.positive_integral_isoton[OF `G \<up> g` G_borel] |
|
1403 |
have *: "(\<lambda>i. T.positive_integral (G i)) \<up> T.positive_integral g" . |
|
1404 |
{ fix i |
|
1405 |
have [simp]: "finite (G i ` space M)" |
|
1406 |
using G(1) unfolding simple_function_def by auto |
|
1407 |
have "T.positive_integral (G i) = T.simple_integral (G i)" |
|
1408 |
using G T.positive_integral_eq_simple_integral by simp |
|
1409 |
also have "\<dots> = positive_integral (\<lambda>x. f x * (\<Sum>y\<in>G i`space M. y * indicator (G i -` {y} \<inter> space M) x))" |
|
1410 |
apply (simp add: T.simple_integral_def) |
|
1411 |
apply (subst positive_integral_cmult[symmetric]) |
|
1412 |
using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) |
|
1413 |
apply (subst positive_integral_setsum[symmetric]) |
|
1414 |
using G_borel assms(1) apply (fastsimp intro: borel_measurable_indicator borel_measurable_vimage) |
|
1415 |
by (simp add: setsum_right_distrib field_simps) |
|
1416 |
also have "\<dots> = positive_integral (\<lambda>x. f x * G i x)" |
|
1417 |
by (auto intro!: positive_integral_cong |
|
1418 |
simp: indicator_def if_distrib setsum_cases) |
|
1419 |
finally have "T.positive_integral (G i) = positive_integral (\<lambda>x. f x * G i x)" . } |
|
1420 |
with * have eq_Tg: "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> T.positive_integral g" by simp |
|
1421 |
from G(2) have "(\<lambda>i x. f x * G i x) \<up> (\<lambda>x. f x * g x)" |
|
1422 |
unfolding isoton_fun_expand by (auto intro!: isoton_cmult_right) |
|
1423 |
then have "(\<lambda>i. positive_integral (\<lambda>x. f x * G i x)) \<up> positive_integral (\<lambda>x. f x * g x)" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1424 |
using assms(1) G_borel by (auto intro!: positive_integral_isoton borel_measurable_pextreal_times) |
39092 | 1425 |
with eq_Tg show "T.positive_integral g = positive_integral (\<lambda>x. f x * g x)" |
1426 |
unfolding isoton_def by simp |
|
1427 |
qed |
|
1428 |
||
38656 | 1429 |
lemma (in measure_space) positive_integral_null_set: |
40859 | 1430 |
assumes "N \<in> null_sets" shows "positive_integral (\<lambda>x. u x * indicator N x) = 0" |
38656 | 1431 |
proof - |
40859 | 1432 |
have "positive_integral (\<lambda>x. u x * indicator N x) = positive_integral (\<lambda>x. 0)" |
1433 |
proof (intro positive_integral_cong_AE AE_I) |
|
1434 |
show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N" |
|
1435 |
by (auto simp: indicator_def) |
|
1436 |
show "\<mu> N = 0" "N \<in> sets M" |
|
1437 |
using assms by auto |
|
35582 | 1438 |
qed |
40859 | 1439 |
then show ?thesis by simp |
38656 | 1440 |
qed |
35582 | 1441 |
|
38656 | 1442 |
lemma (in measure_space) positive_integral_Markov_inequality: |
1443 |
assumes borel: "u \<in> borel_measurable M" and "A \<in> sets M" and c: "c \<noteq> \<omega>" |
|
1444 |
shows "\<mu> ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * positive_integral (\<lambda>x. u x * indicator A x)" |
|
1445 |
(is "\<mu> ?A \<le> _ * ?PI") |
|
1446 |
proof - |
|
1447 |
have "?A \<in> sets M" |
|
1448 |
using `A \<in> sets M` borel by auto |
|
1449 |
hence "\<mu> ?A = positive_integral (\<lambda>x. indicator ?A x)" |
|
1450 |
using positive_integral_indicator by simp |
|
1451 |
also have "\<dots> \<le> positive_integral (\<lambda>x. c * (u x * indicator A x))" |
|
1452 |
proof (rule positive_integral_mono) |
|
1453 |
fix x assume "x \<in> space M" |
|
1454 |
show "indicator ?A x \<le> c * (u x * indicator A x)" |
|
1455 |
by (cases "x \<in> ?A") auto |
|
1456 |
qed |
|
1457 |
also have "\<dots> = c * positive_integral (\<lambda>x. u x * indicator A x)" |
|
1458 |
using assms |
|
1459 |
by (auto intro!: positive_integral_cmult borel_measurable_indicator) |
|
1460 |
finally show ?thesis . |
|
35582 | 1461 |
qed |
1462 |
||
38656 | 1463 |
lemma (in measure_space) positive_integral_0_iff: |
1464 |
assumes borel: "u \<in> borel_measurable M" |
|
1465 |
shows "positive_integral u = 0 \<longleftrightarrow> \<mu> {x\<in>space M. u x \<noteq> 0} = 0" |
|
1466 |
(is "_ \<longleftrightarrow> \<mu> ?A = 0") |
|
35582 | 1467 |
proof - |
38656 | 1468 |
have A: "?A \<in> sets M" using borel by auto |
1469 |
have u: "positive_integral (\<lambda>x. u x * indicator ?A x) = positive_integral u" |
|
1470 |
by (auto intro!: positive_integral_cong simp: indicator_def) |
|
35582 | 1471 |
|
38656 | 1472 |
show ?thesis |
1473 |
proof |
|
1474 |
assume "\<mu> ?A = 0" |
|
1475 |
hence "?A \<in> null_sets" using `?A \<in> sets M` by auto |
|
40859 | 1476 |
from positive_integral_null_set[OF this] |
38656 | 1477 |
have "0 = positive_integral (\<lambda>x. u x * indicator ?A x)" by simp |
1478 |
thus "positive_integral u = 0" unfolding u by simp |
|
1479 |
next |
|
1480 |
assume *: "positive_integral u = 0" |
|
1481 |
let "?M n" = "{x \<in> space M. 1 \<le> of_nat n * u x}" |
|
1482 |
have "0 = (SUP n. \<mu> (?M n \<inter> ?A))" |
|
1483 |
proof - |
|
1484 |
{ fix n |
|
1485 |
from positive_integral_Markov_inequality[OF borel `?A \<in> sets M`, of "of_nat n"] |
|
1486 |
have "\<mu> (?M n \<inter> ?A) = 0" unfolding * u by simp } |
|
1487 |
thus ?thesis by simp |
|
35582 | 1488 |
qed |
38656 | 1489 |
also have "\<dots> = \<mu> (\<Union>n. ?M n \<inter> ?A)" |
1490 |
proof (safe intro!: continuity_from_below) |
|
1491 |
fix n show "?M n \<inter> ?A \<in> sets M" |
|
1492 |
using borel by (auto intro!: Int) |
|
1493 |
next |
|
1494 |
fix n x assume "1 \<le> of_nat n * u x" |
|
1495 |
also have "\<dots> \<le> of_nat (Suc n) * u x" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1496 |
by (subst (1 2) mult_commute) (auto intro!: pextreal_mult_cancel) |
38656 | 1497 |
finally show "1 \<le> of_nat (Suc n) * u x" . |
1498 |
qed |
|
1499 |
also have "\<dots> = \<mu> ?A" |
|
1500 |
proof (safe intro!: arg_cong[where f="\<mu>"]) |
|
1501 |
fix x assume "u x \<noteq> 0" and [simp, intro]: "x \<in> space M" |
|
1502 |
show "x \<in> (\<Union>n. ?M n \<inter> ?A)" |
|
1503 |
proof (cases "u x") |
|
1504 |
case (preal r) |
|
1505 |
obtain j where "1 / r \<le> of_nat j" using ex_le_of_nat .. |
|
1506 |
hence "1 / r * r \<le> of_nat j * r" using preal unfolding mult_le_cancel_right by auto |
|
1507 |
hence "1 \<le> of_nat j * r" using preal `u x \<noteq> 0` by auto |
|
1508 |
thus ?thesis using `u x \<noteq> 0` preal by (auto simp: real_of_nat_def[symmetric]) |
|
1509 |
qed auto |
|
1510 |
qed |
|
1511 |
finally show "\<mu> ?A = 0" by simp |
|
35582 | 1512 |
qed |
1513 |
qed |
|
1514 |
||
39092 | 1515 |
lemma (in measure_space) positive_integral_restricted: |
1516 |
assumes "A \<in> sets M" |
|
1517 |
shows "measure_space.positive_integral (restricted_space A) \<mu> f = positive_integral (\<lambda>x. f x * indicator A x)" |
|
1518 |
(is "measure_space.positive_integral ?R \<mu> f = positive_integral ?f") |
|
1519 |
proof - |
|
1520 |
have msR: "measure_space ?R \<mu>" by (rule restricted_measure_space[OF `A \<in> sets M`]) |
|
1521 |
then interpret R: measure_space ?R \<mu> . |
|
1522 |
have saR: "sigma_algebra ?R" by fact |
|
1523 |
have *: "R.positive_integral f = R.positive_integral ?f" |
|
40859 | 1524 |
by (intro R.positive_integral_cong) auto |
39092 | 1525 |
show ?thesis |
1526 |
unfolding * R.positive_integral_def positive_integral_def |
|
1527 |
unfolding simple_function_restricted[OF `A \<in> sets M`] |
|
1528 |
apply (simp add: SUPR_def) |
|
1529 |
apply (rule arg_cong[where f=Sup]) |
|
40859 | 1530 |
proof (auto simp add: image_iff simple_integral_restricted[OF `A \<in> sets M`]) |
39092 | 1531 |
fix g assume "simple_function (\<lambda>x. g x * indicator A x)" |
40873 | 1532 |
"g \<le> f" |
1533 |
then show "\<exists>x. simple_function x \<and> x \<le> (\<lambda>x. f x * indicator A x) \<and> |
|
39092 | 1534 |
simple_integral (\<lambda>x. g x * indicator A x) = simple_integral x" |
1535 |
apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"]) |
|
1536 |
by (auto simp: indicator_def le_fun_def) |
|
1537 |
next |
|
1538 |
fix g assume g: "simple_function g" "g \<le> (\<lambda>x. f x * indicator A x)" |
|
1539 |
then have *: "(\<lambda>x. g x * indicator A x) = g" |
|
1540 |
"\<And>x. g x * indicator A x = g x" |
|
1541 |
"\<And>x. g x \<le> f x" |
|
39302
d7728f65b353
renamed lemmas: ext_iff -> fun_eq_iff, set_ext_iff -> set_eq_iff, set_ext -> set_eqI
nipkow
parents:
39198
diff
changeset
|
1542 |
by (auto simp: le_fun_def fun_eq_iff indicator_def split: split_if_asm) |
40873 | 1543 |
from g show "\<exists>x. simple_function (\<lambda>xa. x xa * indicator A xa) \<and> x \<le> f \<and> |
39092 | 1544 |
simple_integral g = simple_integral (\<lambda>xa. x xa * indicator A xa)" |
1545 |
using `A \<in> sets M`[THEN sets_into_space] |
|
1546 |
apply (rule_tac exI[of _ "\<lambda>x. g x * indicator A x"]) |
|
1547 |
by (fastsimp simp: le_fun_def *) |
|
1548 |
qed |
|
1549 |
qed |
|
1550 |
||
1551 |
lemma (in measure_space) positive_integral_subalgebra[simp]: |
|
1552 |
assumes borel: "f \<in> borel_measurable (M\<lparr>sets := N\<rparr>)" |
|
1553 |
and N_subalgebra: "N \<subseteq> sets M" "sigma_algebra (M\<lparr>sets := N\<rparr>)" |
|
1554 |
shows "measure_space.positive_integral (M\<lparr>sets := N\<rparr>) \<mu> f = positive_integral f" |
|
1555 |
proof - |
|
1556 |
note msN = measure_space_subalgebra[OF N_subalgebra] |
|
1557 |
then interpret N: measure_space "M\<lparr>sets:=N\<rparr>" \<mu> . |
|
1558 |
from N.borel_measurable_implies_simple_function_sequence[OF borel] |
|
1559 |
obtain fs where Nsf: "\<And>i. N.simple_function (fs i)" and "fs \<up> f" by blast |
|
1560 |
then have sf: "\<And>i. simple_function (fs i)" |
|
1561 |
using simple_function_subalgebra[OF _ N_subalgebra] by blast |
|
1562 |
from positive_integral_isoton_simple[OF `fs \<up> f` sf] N.positive_integral_isoton_simple[OF `fs \<up> f` Nsf] |
|
1563 |
show ?thesis unfolding simple_integral_subalgebra[OF msN] isoton_def by simp |
|
1564 |
qed |
|
1565 |
||
35692 | 1566 |
section "Lebesgue Integral" |
1567 |
||
38656 | 1568 |
definition (in measure_space) integrable where |
1569 |
"integrable f \<longleftrightarrow> f \<in> borel_measurable M \<and> |
|
1570 |
positive_integral (\<lambda>x. Real (f x)) \<noteq> \<omega> \<and> |
|
1571 |
positive_integral (\<lambda>x. Real (- f x)) \<noteq> \<omega>" |
|
35692 | 1572 |
|
38656 | 1573 |
lemma (in measure_space) integrableD[dest]: |
1574 |
assumes "integrable f" |
|
1575 |
shows "f \<in> borel_measurable M" |
|
1576 |
"positive_integral (\<lambda>x. Real (f x)) \<noteq> \<omega>" |
|
1577 |
"positive_integral (\<lambda>x. Real (- f x)) \<noteq> \<omega>" |
|
1578 |
using assms unfolding integrable_def by auto |
|
35692 | 1579 |
|
38656 | 1580 |
definition (in measure_space) integral where |
1581 |
"integral f = |
|
1582 |
real (positive_integral (\<lambda>x. Real (f x))) - |
|
1583 |
real (positive_integral (\<lambda>x. Real (- f x)))" |
|
1584 |
||
1585 |
lemma (in measure_space) integral_cong: |
|
35582 | 1586 |
assumes cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x" |
1587 |
shows "integral f = integral g" |
|
38656 | 1588 |
using assms by (simp cong: positive_integral_cong add: integral_def) |
35582 | 1589 |
|
40859 | 1590 |
lemma (in measure_space) integral_cong_measure: |
1591 |
assumes "\<And>A. A \<in> sets M \<Longrightarrow> \<nu> A = \<mu> A" |
|
1592 |
shows "measure_space.integral M \<nu> f = integral f" |
|
1593 |
proof - |
|
1594 |
interpret v: measure_space M \<nu> |
|
1595 |
by (rule measure_space_cong) fact |
|
1596 |
show ?thesis |
|
1597 |
unfolding integral_def v.integral_def |
|
1598 |
by (simp add: positive_integral_cong_measure[OF assms]) |
|
1599 |
qed |
|
1600 |
||
1601 |
lemma (in measure_space) integral_cong_AE: |
|
1602 |
assumes cong: "AE x. f x = g x" |
|
1603 |
shows "integral f = integral g" |
|
1604 |
proof - |
|
1605 |
have "AE x. Real (f x) = Real (g x)" |
|
1606 |
using cong by (rule AE_mp) simp |
|
1607 |
moreover have "AE x. Real (- f x) = Real (- g x)" |
|
1608 |
using cong by (rule AE_mp) simp |
|
1609 |
ultimately show ?thesis |
|
1610 |
by (simp cong: positive_integral_cong_AE add: integral_def) |
|
1611 |
qed |
|
1612 |
||
38656 | 1613 |
lemma (in measure_space) integrable_cong: |
1614 |
"(\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> integrable f \<longleftrightarrow> integrable g" |
|
1615 |
by (simp cong: positive_integral_cong measurable_cong add: integrable_def) |
|
1616 |
||
1617 |
lemma (in measure_space) integral_eq_positive_integral: |
|
1618 |
assumes "\<And>x. 0 \<le> f x" |
|
1619 |
shows "integral f = real (positive_integral (\<lambda>x. Real (f x)))" |
|
35582 | 1620 |
proof - |
38656 | 1621 |
have "\<And>x. Real (- f x) = 0" using assms by simp |
1622 |
thus ?thesis by (simp del: Real_eq_0 add: integral_def) |
|
35582 | 1623 |
qed |
1624 |
||
40859 | 1625 |
lemma (in measure_space) integral_vimage_inv: |
1626 |
assumes f: "bij_betw f S (space M)" |
|
1627 |
shows "measure_space.integral (vimage_algebra S f) (\<lambda>A. \<mu> (f ` A)) (\<lambda>x. g x) = integral (\<lambda>x. g (the_inv_into S f x))" |
|
1628 |
proof - |
|
1629 |
interpret v: measure_space "vimage_algebra S f" "\<lambda>A. \<mu> (f ` A)" |
|
1630 |
using f by (rule measure_space_isomorphic) |
|
1631 |
have "\<And>x. x \<in> space (vimage_algebra S f) \<Longrightarrow> the_inv_into S f (f x) = x" |
|
1632 |
using f[unfolded bij_betw_def] by (simp add: the_inv_into_f_f) |
|
1633 |
then have *: "v.positive_integral (\<lambda>x. Real (g (the_inv_into S f (f x)))) = v.positive_integral (\<lambda>x. Real (g x))" |
|
1634 |
"v.positive_integral (\<lambda>x. Real (- g (the_inv_into S f (f x)))) = v.positive_integral (\<lambda>x. Real (- g x))" |
|
1635 |
by (auto intro!: v.positive_integral_cong) |
|
1636 |
show ?thesis |
|
1637 |
unfolding integral_def v.integral_def |
|
1638 |
unfolding positive_integral_vimage[OF f] |
|
1639 |
by (simp add: *) |
|
1640 |
qed |
|
1641 |
||
38656 | 1642 |
lemma (in measure_space) integral_minus[intro, simp]: |
1643 |
assumes "integrable f" |
|
1644 |
shows "integrable (\<lambda>x. - f x)" "integral (\<lambda>x. - f x) = - integral f" |
|
1645 |
using assms by (auto simp: integrable_def integral_def) |
|
1646 |
||
1647 |
lemma (in measure_space) integral_of_positive_diff: |
|
1648 |
assumes integrable: "integrable u" "integrable v" |
|
1649 |
and f_def: "\<And>x. f x = u x - v x" and pos: "\<And>x. 0 \<le> u x" "\<And>x. 0 \<le> v x" |
|
1650 |
shows "integrable f" and "integral f = integral u - integral v" |
|
35582 | 1651 |
proof - |
38656 | 1652 |
let ?PI = positive_integral |
1653 |
let "?f x" = "Real (f x)" |
|
1654 |
let "?mf x" = "Real (- f x)" |
|
1655 |
let "?u x" = "Real (u x)" |
|
1656 |
let "?v x" = "Real (v x)" |
|
1657 |
||
1658 |
from borel_measurable_diff[of u v] integrable |
|
1659 |
have f_borel: "?f \<in> borel_measurable M" and |
|
1660 |
mf_borel: "?mf \<in> borel_measurable M" and |
|
1661 |
v_borel: "?v \<in> borel_measurable M" and |
|
1662 |
u_borel: "?u \<in> borel_measurable M" and |
|
1663 |
"f \<in> borel_measurable M" |
|
1664 |
by (auto simp: f_def[symmetric] integrable_def) |
|
35582 | 1665 |
|
38656 | 1666 |
have "?PI (\<lambda>x. Real (u x - v x)) \<le> ?PI ?u" |
1667 |
using pos by (auto intro!: positive_integral_mono) |
|
1668 |
moreover have "?PI (\<lambda>x. Real (v x - u x)) \<le> ?PI ?v" |
|
1669 |
using pos by (auto intro!: positive_integral_mono) |
|
1670 |
ultimately show f: "integrable f" |
|
1671 |
using `integrable u` `integrable v` `f \<in> borel_measurable M` |
|
1672 |
by (auto simp: integrable_def f_def) |
|
1673 |
hence mf: "integrable (\<lambda>x. - f x)" .. |
|
1674 |
||
1675 |
have *: "\<And>x. Real (- v x) = 0" "\<And>x. Real (- u x) = 0" |
|
1676 |
using pos by auto |
|
35582 | 1677 |
|
38656 | 1678 |
have "\<And>x. ?u x + ?mf x = ?v x + ?f x" |
1679 |
unfolding f_def using pos by simp |
|
1680 |
hence "?PI (\<lambda>x. ?u x + ?mf x) = ?PI (\<lambda>x. ?v x + ?f x)" by simp |
|
1681 |
hence "real (?PI ?u + ?PI ?mf) = real (?PI ?v + ?PI ?f)" |
|
1682 |
using positive_integral_add[OF u_borel mf_borel] |
|
1683 |
using positive_integral_add[OF v_borel f_borel] |
|
1684 |
by auto |
|
1685 |
then show "integral f = integral u - integral v" |
|
1686 |
using f mf `integrable u` `integrable v` |
|
1687 |
unfolding integral_def integrable_def * |
|
1688 |
by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?v", cases "?PI ?u") |
|
1689 |
(auto simp add: field_simps) |
|
35582 | 1690 |
qed |
1691 |
||
38656 | 1692 |
lemma (in measure_space) integral_linear: |
1693 |
assumes "integrable f" "integrable g" and "0 \<le> a" |
|
1694 |
shows "integrable (\<lambda>t. a * f t + g t)" |
|
1695 |
and "integral (\<lambda>t. a * f t + g t) = a * integral f + integral g" |
|
1696 |
proof - |
|
1697 |
let ?PI = positive_integral |
|
1698 |
let "?f x" = "Real (f x)" |
|
1699 |
let "?g x" = "Real (g x)" |
|
1700 |
let "?mf x" = "Real (- f x)" |
|
1701 |
let "?mg x" = "Real (- g x)" |
|
1702 |
let "?p t" = "max 0 (a * f t) + max 0 (g t)" |
|
1703 |
let "?n t" = "max 0 (- (a * f t)) + max 0 (- g t)" |
|
1704 |
||
1705 |
have pos: "?f \<in> borel_measurable M" "?g \<in> borel_measurable M" |
|
1706 |
and neg: "?mf \<in> borel_measurable M" "?mg \<in> borel_measurable M" |
|
1707 |
and p: "?p \<in> borel_measurable M" |
|
1708 |
and n: "?n \<in> borel_measurable M" |
|
1709 |
using assms by (simp_all add: integrable_def) |
|
35582 | 1710 |
|
38656 | 1711 |
have *: "\<And>x. Real (?p x) = Real a * ?f x + ?g x" |
1712 |
"\<And>x. Real (?n x) = Real a * ?mf x + ?mg x" |
|
1713 |
"\<And>x. Real (- ?p x) = 0" |
|
1714 |
"\<And>x. Real (- ?n x) = 0" |
|
1715 |
using `0 \<le> a` by (auto simp: max_def min_def zero_le_mult_iff mult_le_0_iff add_nonpos_nonpos) |
|
1716 |
||
1717 |
note linear = |
|
1718 |
positive_integral_linear[OF pos] |
|
1719 |
positive_integral_linear[OF neg] |
|
35582 | 1720 |
|
38656 | 1721 |
have "integrable ?p" "integrable ?n" |
1722 |
"\<And>t. a * f t + g t = ?p t - ?n t" "\<And>t. 0 \<le> ?p t" "\<And>t. 0 \<le> ?n t" |
|
1723 |
using assms p n unfolding integrable_def * linear by auto |
|
1724 |
note diff = integral_of_positive_diff[OF this] |
|
1725 |
||
1726 |
show "integrable (\<lambda>t. a * f t + g t)" by (rule diff) |
|
1727 |
||
1728 |
from assms show "integral (\<lambda>t. a * f t + g t) = a * integral f + integral g" |
|
1729 |
unfolding diff(2) unfolding integral_def * linear integrable_def |
|
1730 |
by (cases "?PI ?f", cases "?PI ?mf", cases "?PI ?g", cases "?PI ?mg") |
|
1731 |
(auto simp add: field_simps zero_le_mult_iff) |
|
1732 |
qed |
|
1733 |
||
1734 |
lemma (in measure_space) integral_add[simp, intro]: |
|
1735 |
assumes "integrable f" "integrable g" |
|
35582 | 1736 |
shows "integrable (\<lambda>t. f t + g t)" |
1737 |
and "integral (\<lambda>t. f t + g t) = integral f + integral g" |
|
38656 | 1738 |
using assms integral_linear[where a=1] by auto |
1739 |
||
1740 |
lemma (in measure_space) integral_zero[simp, intro]: |
|
1741 |
shows "integrable (\<lambda>x. 0)" |
|
1742 |
and "integral (\<lambda>x. 0) = 0" |
|
1743 |
unfolding integrable_def integral_def |
|
1744 |
by (auto simp add: borel_measurable_const) |
|
35582 | 1745 |
|
38656 | 1746 |
lemma (in measure_space) integral_cmult[simp, intro]: |
1747 |
assumes "integrable f" |
|
1748 |
shows "integrable (\<lambda>t. a * f t)" (is ?P) |
|
1749 |
and "integral (\<lambda>t. a * f t) = a * integral f" (is ?I) |
|
1750 |
proof - |
|
1751 |
have "integrable (\<lambda>t. a * f t) \<and> integral (\<lambda>t. a * f t) = a * integral f" |
|
1752 |
proof (cases rule: le_cases) |
|
1753 |
assume "0 \<le> a" show ?thesis |
|
1754 |
using integral_linear[OF assms integral_zero(1) `0 \<le> a`] |
|
1755 |
by (simp add: integral_zero) |
|
1756 |
next |
|
1757 |
assume "a \<le> 0" hence "0 \<le> - a" by auto |
|
1758 |
have *: "\<And>t. - a * t + 0 = (-a) * t" by simp |
|
1759 |
show ?thesis using integral_linear[OF assms integral_zero(1) `0 \<le> - a`] |
|
1760 |
integral_minus(1)[of "\<lambda>t. - a * f t"] |
|
1761 |
unfolding * integral_zero by simp |
|
1762 |
qed |
|
1763 |
thus ?P ?I by auto |
|
35582 | 1764 |
qed |
1765 |
||
40859 | 1766 |
lemma (in measure_space) integral_mono_AE: |
1767 |
assumes fg: "integrable f" "integrable g" |
|
1768 |
and mono: "AE t. f t \<le> g t" |
|
1769 |
shows "integral f \<le> integral g" |
|
1770 |
proof - |
|
1771 |
have "AE x. Real (f x) \<le> Real (g x)" |
|
1772 |
using mono by (rule AE_mp) (auto intro!: AE_cong) |
|
1773 |
moreover have "AE x. Real (- g x) \<le> Real (- f x)" |
|
1774 |
using mono by (rule AE_mp) (auto intro!: AE_cong) |
|
1775 |
ultimately show ?thesis using fg |
|
1776 |
by (auto simp: integral_def integrable_def diff_minus |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1777 |
intro!: add_mono real_of_pextreal_mono positive_integral_mono_AE) |
40859 | 1778 |
qed |
1779 |
||
38656 | 1780 |
lemma (in measure_space) integral_mono: |
1781 |
assumes fg: "integrable f" "integrable g" |
|
35582 | 1782 |
and mono: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t" |
1783 |
shows "integral f \<le> integral g" |
|
40859 | 1784 |
apply (rule integral_mono_AE[OF fg]) |
1785 |
using mono by (rule AE_cong) auto |
|
35582 | 1786 |
|
38656 | 1787 |
lemma (in measure_space) integral_diff[simp, intro]: |
1788 |
assumes f: "integrable f" and g: "integrable g" |
|
1789 |
shows "integrable (\<lambda>t. f t - g t)" |
|
1790 |
and "integral (\<lambda>t. f t - g t) = integral f - integral g" |
|
1791 |
using integral_add[OF f integral_minus(1)[OF g]] |
|
1792 |
unfolding diff_minus integral_minus(2)[OF g] |
|
1793 |
by auto |
|
1794 |
||
1795 |
lemma (in measure_space) integral_indicator[simp, intro]: |
|
1796 |
assumes "a \<in> sets M" and "\<mu> a \<noteq> \<omega>" |
|
1797 |
shows "integral (indicator a) = real (\<mu> a)" (is ?int) |
|
1798 |
and "integrable (indicator a)" (is ?able) |
|
35582 | 1799 |
proof - |
38656 | 1800 |
have *: |
1801 |
"\<And>A x. Real (indicator A x) = indicator A x" |
|
1802 |
"\<And>A x. Real (- indicator A x) = 0" unfolding indicator_def by auto |
|
1803 |
show ?int ?able |
|
1804 |
using assms unfolding integral_def integrable_def |
|
1805 |
by (auto simp: * positive_integral_indicator borel_measurable_indicator) |
|
35582 | 1806 |
qed |
1807 |
||
38656 | 1808 |
lemma (in measure_space) integral_cmul_indicator: |
1809 |
assumes "A \<in> sets M" and "c \<noteq> 0 \<Longrightarrow> \<mu> A \<noteq> \<omega>" |
|
1810 |
shows "integrable (\<lambda>x. c * indicator A x)" (is ?P) |
|
1811 |
and "integral (\<lambda>x. c * indicator A x) = c * real (\<mu> A)" (is ?I) |
|
1812 |
proof - |
|
1813 |
show ?P |
|
1814 |
proof (cases "c = 0") |
|
1815 |
case False with assms show ?thesis by simp |
|
1816 |
qed simp |
|
35582 | 1817 |
|
38656 | 1818 |
show ?I |
1819 |
proof (cases "c = 0") |
|
1820 |
case False with assms show ?thesis by simp |
|
1821 |
qed simp |
|
1822 |
qed |
|
35582 | 1823 |
|
38656 | 1824 |
lemma (in measure_space) integral_setsum[simp, intro]: |
35582 | 1825 |
assumes "\<And>n. n \<in> S \<Longrightarrow> integrable (f n)" |
1826 |
shows "integral (\<lambda>x. \<Sum> i \<in> S. f i x) = (\<Sum> i \<in> S. integral (f i))" (is "?int S") |
|
38656 | 1827 |
and "integrable (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I S") |
35582 | 1828 |
proof - |
38656 | 1829 |
have "?int S \<and> ?I S" |
1830 |
proof (cases "finite S") |
|
1831 |
assume "finite S" |
|
1832 |
from this assms show ?thesis by (induct S) simp_all |
|
1833 |
qed simp |
|
35582 | 1834 |
thus "?int S" and "?I S" by auto |
1835 |
qed |
|
1836 |
||
36624 | 1837 |
lemma (in measure_space) integrable_abs: |
1838 |
assumes "integrable f" |
|
1839 |
shows "integrable (\<lambda> x. \<bar>f x\<bar>)" |
|
1840 |
proof - |
|
38656 | 1841 |
have *: |
1842 |
"\<And>x. Real \<bar>f x\<bar> = Real (f x) + Real (- f x)" |
|
1843 |
"\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto |
|
1844 |
have abs: "(\<lambda>x. \<bar>f x\<bar>) \<in> borel_measurable M" and |
|
1845 |
f: "(\<lambda>x. Real (f x)) \<in> borel_measurable M" |
|
1846 |
"(\<lambda>x. Real (- f x)) \<in> borel_measurable M" |
|
1847 |
using assms unfolding integrable_def by auto |
|
1848 |
from abs assms show ?thesis unfolding integrable_def * |
|
1849 |
using positive_integral_linear[OF f, of 1] by simp |
|
1850 |
qed |
|
1851 |
||
1852 |
lemma (in measure_space) integrable_bound: |
|
1853 |
assumes "integrable f" |
|
1854 |
and f: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" |
|
1855 |
"\<And>x. x \<in> space M \<Longrightarrow> \<bar>g x\<bar> \<le> f x" |
|
1856 |
assumes borel: "g \<in> borel_measurable M" |
|
1857 |
shows "integrable g" |
|
1858 |
proof - |
|
1859 |
have "positive_integral (\<lambda>x. Real (g x)) \<le> positive_integral (\<lambda>x. Real \<bar>g x\<bar>)" |
|
1860 |
by (auto intro!: positive_integral_mono) |
|
1861 |
also have "\<dots> \<le> positive_integral (\<lambda>x. Real (f x))" |
|
1862 |
using f by (auto intro!: positive_integral_mono) |
|
1863 |
also have "\<dots> < \<omega>" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1864 |
using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>) |
38656 | 1865 |
finally have pos: "positive_integral (\<lambda>x. Real (g x)) < \<omega>" . |
1866 |
||
1867 |
have "positive_integral (\<lambda>x. Real (- g x)) \<le> positive_integral (\<lambda>x. Real (\<bar>g x\<bar>))" |
|
1868 |
by (auto intro!: positive_integral_mono) |
|
1869 |
also have "\<dots> \<le> positive_integral (\<lambda>x. Real (f x))" |
|
1870 |
using f by (auto intro!: positive_integral_mono) |
|
1871 |
also have "\<dots> < \<omega>" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
1872 |
using `integrable f` unfolding integrable_def by (auto simp: pextreal_less_\<omega>) |
38656 | 1873 |
finally have neg: "positive_integral (\<lambda>x. Real (- g x)) < \<omega>" . |
1874 |
||
1875 |
from neg pos borel show ?thesis |
|
36624 | 1876 |
unfolding integrable_def by auto |
38656 | 1877 |
qed |
1878 |
||
1879 |
lemma (in measure_space) integrable_abs_iff: |
|
1880 |
"f \<in> borel_measurable M \<Longrightarrow> integrable (\<lambda> x. \<bar>f x\<bar>) \<longleftrightarrow> integrable f" |
|
1881 |
by (auto intro!: integrable_bound[where g=f] integrable_abs) |
|
1882 |
||
1883 |
lemma (in measure_space) integrable_max: |
|
1884 |
assumes int: "integrable f" "integrable g" |
|
1885 |
shows "integrable (\<lambda> x. max (f x) (g x))" |
|
1886 |
proof (rule integrable_bound) |
|
1887 |
show "integrable (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)" |
|
1888 |
using int by (simp add: integrable_abs) |
|
1889 |
show "(\<lambda>x. max (f x) (g x)) \<in> borel_measurable M" |
|
1890 |
using int unfolding integrable_def by auto |
|
1891 |
next |
|
1892 |
fix x assume "x \<in> space M" |
|
1893 |
show "0 \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" "\<bar>max (f x) (g x)\<bar> \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" |
|
1894 |
by auto |
|
1895 |
qed |
|
1896 |
||
1897 |
lemma (in measure_space) integrable_min: |
|
1898 |
assumes int: "integrable f" "integrable g" |
|
1899 |
shows "integrable (\<lambda> x. min (f x) (g x))" |
|
1900 |
proof (rule integrable_bound) |
|
1901 |
show "integrable (\<lambda>x. \<bar>f x\<bar> + \<bar>g x\<bar>)" |
|
1902 |
using int by (simp add: integrable_abs) |
|
1903 |
show "(\<lambda>x. min (f x) (g x)) \<in> borel_measurable M" |
|
1904 |
using int unfolding integrable_def by auto |
|
1905 |
next |
|
1906 |
fix x assume "x \<in> space M" |
|
1907 |
show "0 \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" "\<bar>min (f x) (g x)\<bar> \<le> \<bar>f x\<bar> + \<bar>g x\<bar>" |
|
1908 |
by auto |
|
1909 |
qed |
|
1910 |
||
1911 |
lemma (in measure_space) integral_triangle_inequality: |
|
1912 |
assumes "integrable f" |
|
1913 |
shows "\<bar>integral f\<bar> \<le> integral (\<lambda>x. \<bar>f x\<bar>)" |
|
1914 |
proof - |
|
1915 |
have "\<bar>integral f\<bar> = max (integral f) (- integral f)" by auto |
|
1916 |
also have "\<dots> \<le> integral (\<lambda>x. \<bar>f x\<bar>)" |
|
1917 |
using assms integral_minus(2)[of f, symmetric] |
|
1918 |
by (auto intro!: integral_mono integrable_abs simp del: integral_minus) |
|
1919 |
finally show ?thesis . |
|
36624 | 1920 |
qed |
1921 |
||
38656 | 1922 |
lemma (in measure_space) integral_positive: |
1923 |
assumes "integrable f" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" |
|
1924 |
shows "0 \<le> integral f" |
|
1925 |
proof - |
|
1926 |
have "0 = integral (\<lambda>x. 0)" by (auto simp: integral_zero) |
|
1927 |
also have "\<dots> \<le> integral f" |
|
1928 |
using assms by (rule integral_mono[OF integral_zero(1)]) |
|
1929 |
finally show ?thesis . |
|
1930 |
qed |
|
1931 |
||
1932 |
lemma (in measure_space) integral_monotone_convergence_pos: |
|
1933 |
assumes i: "\<And>i. integrable (f i)" and mono: "\<And>x. mono (\<lambda>n. f n x)" |
|
1934 |
and pos: "\<And>x i. 0 \<le> f i x" |
|
1935 |
and lim: "\<And>x. (\<lambda>i. f i x) ----> u x" |
|
1936 |
and ilim: "(\<lambda>i. integral (f i)) ----> x" |
|
1937 |
shows "integrable u" |
|
1938 |
and "integral u = x" |
|
35582 | 1939 |
proof - |
38656 | 1940 |
{ fix x have "0 \<le> u x" |
1941 |
using mono pos[of 0 x] incseq_le[OF _ lim, of x 0] |
|
1942 |
by (simp add: mono_def incseq_def) } |
|
1943 |
note pos_u = this |
|
1944 |
||
1945 |
hence [simp]: "\<And>i x. Real (- f i x) = 0" "\<And>x. Real (- u x) = 0" |
|
1946 |
using pos by auto |
|
1947 |
||
1948 |
have SUP_F: "\<And>x. (SUP n. Real (f n x)) = Real (u x)" |
|
1949 |
using mono pos pos_u lim by (rule SUP_eq_LIMSEQ[THEN iffD2]) |
|
1950 |
||
1951 |
have borel_f: "\<And>i. (\<lambda>x. Real (f i x)) \<in> borel_measurable M" |
|
1952 |
using i unfolding integrable_def by auto |
|
1953 |
hence "(SUP i. (\<lambda>x. Real (f i x))) \<in> borel_measurable M" |
|
35582 | 1954 |
by auto |
38656 | 1955 |
hence borel_u: "u \<in> borel_measurable M" |
1956 |
using pos_u by (auto simp: borel_measurable_Real_eq SUPR_fun_expand SUP_F) |
|
1957 |
||
1958 |
have integral_eq: "\<And>n. positive_integral (\<lambda>x. Real (f n x)) = Real (integral (f n))" |
|
1959 |
using i unfolding integral_def integrable_def by (auto simp: Real_real) |
|
1960 |
||
1961 |
have pos_integral: "\<And>n. 0 \<le> integral (f n)" |
|
1962 |
using pos i by (auto simp: integral_positive) |
|
1963 |
hence "0 \<le> x" |
|
1964 |
using LIMSEQ_le_const[OF ilim, of 0] by auto |
|
1965 |
||
1966 |
have "(\<lambda>i. positive_integral (\<lambda>x. Real (f i x))) \<up> positive_integral (\<lambda>x. Real (u x))" |
|
1967 |
proof (rule positive_integral_isoton) |
|
1968 |
from SUP_F mono pos |
|
1969 |
show "(\<lambda>i x. Real (f i x)) \<up> (\<lambda>x. Real (u x))" |
|
1970 |
unfolding isoton_fun_expand by (auto simp: isoton_def mono_def) |
|
1971 |
qed (rule borel_f) |
|
1972 |
hence pI: "positive_integral (\<lambda>x. Real (u x)) = |
|
1973 |
(SUP n. positive_integral (\<lambda>x. Real (f n x)))" |
|
1974 |
unfolding isoton_def by simp |
|
1975 |
also have "\<dots> = Real x" unfolding integral_eq |
|
1976 |
proof (rule SUP_eq_LIMSEQ[THEN iffD2]) |
|
1977 |
show "mono (\<lambda>n. integral (f n))" |
|
1978 |
using mono i by (auto simp: mono_def intro!: integral_mono) |
|
1979 |
show "\<And>n. 0 \<le> integral (f n)" using pos_integral . |
|
1980 |
show "0 \<le> x" using `0 \<le> x` . |
|
1981 |
show "(\<lambda>n. integral (f n)) ----> x" using ilim . |
|
1982 |
qed |
|
1983 |
finally show "integrable u" "integral u = x" using borel_u `0 \<le> x` |
|
1984 |
unfolding integrable_def integral_def by auto |
|
1985 |
qed |
|
1986 |
||
1987 |
lemma (in measure_space) integral_monotone_convergence: |
|
1988 |
assumes f: "\<And>i. integrable (f i)" and "mono f" |
|
1989 |
and lim: "\<And>x. (\<lambda>i. f i x) ----> u x" |
|
1990 |
and ilim: "(\<lambda>i. integral (f i)) ----> x" |
|
1991 |
shows "integrable u" |
|
1992 |
and "integral u = x" |
|
1993 |
proof - |
|
1994 |
have 1: "\<And>i. integrable (\<lambda>x. f i x - f 0 x)" |
|
1995 |
using f by (auto intro!: integral_diff) |
|
1996 |
have 2: "\<And>x. mono (\<lambda>n. f n x - f 0 x)" using `mono f` |
|
1997 |
unfolding mono_def le_fun_def by auto |
|
1998 |
have 3: "\<And>x n. 0 \<le> f n x - f 0 x" using `mono f` |
|
1999 |
unfolding mono_def le_fun_def by (auto simp: field_simps) |
|
2000 |
have 4: "\<And>x. (\<lambda>i. f i x - f 0 x) ----> u x - f 0 x" |
|
2001 |
using lim by (auto intro!: LIMSEQ_diff) |
|
2002 |
have 5: "(\<lambda>i. integral (\<lambda>x. f i x - f 0 x)) ----> x - integral (f 0)" |
|
2003 |
using f ilim by (auto intro!: LIMSEQ_diff simp: integral_diff) |
|
2004 |
note diff = integral_monotone_convergence_pos[OF 1 2 3 4 5] |
|
2005 |
have "integrable (\<lambda>x. (u x - f 0 x) + f 0 x)" |
|
2006 |
using diff(1) f by (rule integral_add(1)) |
|
2007 |
with diff(2) f show "integrable u" "integral u = x" |
|
2008 |
by (auto simp: integral_diff) |
|
2009 |
qed |
|
2010 |
||
2011 |
lemma (in measure_space) integral_0_iff: |
|
2012 |
assumes "integrable f" |
|
2013 |
shows "integral (\<lambda>x. \<bar>f x\<bar>) = 0 \<longleftrightarrow> \<mu> {x\<in>space M. f x \<noteq> 0} = 0" |
|
2014 |
proof - |
|
2015 |
have *: "\<And>x. Real (- \<bar>f x\<bar>) = 0" by auto |
|
2016 |
have "integrable (\<lambda>x. \<bar>f x\<bar>)" using assms by (rule integrable_abs) |
|
2017 |
hence "(\<lambda>x. Real (\<bar>f x\<bar>)) \<in> borel_measurable M" |
|
2018 |
"positive_integral (\<lambda>x. Real \<bar>f x\<bar>) \<noteq> \<omega>" unfolding integrable_def by auto |
|
2019 |
from positive_integral_0_iff[OF this(1)] this(2) |
|
2020 |
show ?thesis unfolding integral_def * |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
2021 |
by (simp add: real_of_pextreal_eq_0) |
35582 | 2022 |
qed |
2023 |
||
40859 | 2024 |
lemma (in measure_space) positive_integral_omega: |
2025 |
assumes "f \<in> borel_measurable M" |
|
2026 |
and "positive_integral f \<noteq> \<omega>" |
|
2027 |
shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0" |
|
2028 |
proof - |
|
2029 |
have "\<omega> * \<mu> (f -` {\<omega>} \<inter> space M) = positive_integral (\<lambda>x. \<omega> * indicator (f -` {\<omega>} \<inter> space M) x)" |
|
2030 |
using positive_integral_cmult_indicator[OF borel_measurable_vimage, OF assms(1), of \<omega> \<omega>] by simp |
|
2031 |
also have "\<dots> \<le> positive_integral f" |
|
2032 |
by (auto intro!: positive_integral_mono simp: indicator_def) |
|
2033 |
finally show ?thesis |
|
2034 |
using assms(2) by (cases ?thesis) auto |
|
2035 |
qed |
|
2036 |
||
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2037 |
lemma (in measure_space) positive_integral_omega_AE: |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2038 |
assumes "f \<in> borel_measurable M" "positive_integral f \<noteq> \<omega>" shows "AE x. f x \<noteq> \<omega>" |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2039 |
proof (rule AE_I) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2040 |
show "\<mu> (f -` {\<omega>} \<inter> space M) = 0" |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2041 |
by (rule positive_integral_omega[OF assms]) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2042 |
show "f -` {\<omega>} \<inter> space M \<in> sets M" |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2043 |
using assms by (auto intro: borel_measurable_vimage) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2044 |
qed auto |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2045 |
|
40859 | 2046 |
lemma (in measure_space) simple_integral_omega: |
2047 |
assumes "simple_function f" |
|
2048 |
and "simple_integral f \<noteq> \<omega>" |
|
2049 |
shows "\<mu> (f -` {\<omega>} \<inter> space M) = 0" |
|
2050 |
proof (rule positive_integral_omega) |
|
2051 |
show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) |
|
2052 |
show "positive_integral f \<noteq> \<omega>" |
|
2053 |
using assms by (simp add: positive_integral_eq_simple_integral) |
|
2054 |
qed |
|
2055 |
||
41026
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2056 |
lemma (in measure_space) integral_real: |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2057 |
fixes f :: "'a \<Rightarrow> pextreal" |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2058 |
assumes "AE x. f x \<noteq> \<omega>" |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2059 |
shows "integral (\<lambda>x. real (f x)) = real (positive_integral f)" (is ?plus) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2060 |
and "integral (\<lambda>x. - real (f x)) = - real (positive_integral f)" (is ?minus) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2061 |
proof - |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2062 |
have "positive_integral (\<lambda>x. Real (real (f x))) = positive_integral f" |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2063 |
apply (rule positive_integral_cong_AE) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2064 |
apply (rule AE_mp[OF assms(1)]) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2065 |
by (auto intro!: AE_cong simp: Real_real) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2066 |
moreover |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2067 |
have "positive_integral (\<lambda>x. Real (- real (f x))) = positive_integral (\<lambda>x. 0)" |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2068 |
by (intro positive_integral_cong) auto |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2069 |
ultimately show ?plus ?minus |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2070 |
by (auto simp: integral_def integrable_def) |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2071 |
qed |
bea75746dc9d
folding on arbitrary Lebesgue integrable functions
hoelzl
parents:
41023
diff
changeset
|
2072 |
|
38656 | 2073 |
lemma (in measure_space) integral_dominated_convergence: |
2074 |
assumes u: "\<And>i. integrable (u i)" and bound: "\<And>x j. x\<in>space M \<Longrightarrow> \<bar>u j x\<bar> \<le> w x" |
|
2075 |
and w: "integrable w" "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> w x" |
|
2076 |
and u': "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. u i x) ----> u' x" |
|
2077 |
shows "integrable u'" |
|
2078 |
and "(\<lambda>i. integral (\<lambda>x. \<bar>u i x - u' x\<bar>)) ----> 0" (is "?lim_diff") |
|
2079 |
and "(\<lambda>i. integral (u i)) ----> integral u'" (is ?lim) |
|
36624 | 2080 |
proof - |
38656 | 2081 |
{ fix x j assume x: "x \<in> space M" |
2082 |
from u'[OF x] have "(\<lambda>i. \<bar>u i x\<bar>) ----> \<bar>u' x\<bar>" by (rule LIMSEQ_imp_rabs) |
|
2083 |
from LIMSEQ_le_const2[OF this] |
|
2084 |
have "\<bar>u' x\<bar> \<le> w x" using bound[OF x] by auto } |
|
2085 |
note u'_bound = this |
|
2086 |
||
2087 |
from u[unfolded integrable_def] |
|
2088 |
have u'_borel: "u' \<in> borel_measurable M" |
|
2089 |
using u' by (blast intro: borel_measurable_LIMSEQ[of u]) |
|
2090 |
||
2091 |
show "integrable u'" |
|
2092 |
proof (rule integrable_bound) |
|
2093 |
show "integrable w" by fact |
|
2094 |
show "u' \<in> borel_measurable M" by fact |
|
2095 |
next |
|
2096 |
fix x assume x: "x \<in> space M" |
|
2097 |
thus "0 \<le> w x" by fact |
|
2098 |
show "\<bar>u' x\<bar> \<le> w x" using u'_bound[OF x] . |
|
2099 |
qed |
|
2100 |
||
2101 |
let "?diff n x" = "2 * w x - \<bar>u n x - u' x\<bar>" |
|
2102 |
have diff: "\<And>n. integrable (\<lambda>x. \<bar>u n x - u' x\<bar>)" |
|
2103 |
using w u `integrable u'` |
|
2104 |
by (auto intro!: integral_add integral_diff integral_cmult integrable_abs) |
|
2105 |
||
2106 |
{ fix j x assume x: "x \<in> space M" |
|
2107 |
have "\<bar>u j x - u' x\<bar> \<le> \<bar>u j x\<bar> + \<bar>u' x\<bar>" by auto |
|
2108 |
also have "\<dots> \<le> w x + w x" |
|
2109 |
by (rule add_mono[OF bound[OF x] u'_bound[OF x]]) |
|
2110 |
finally have "\<bar>u j x - u' x\<bar> \<le> 2 * w x" by simp } |
|
2111 |
note diff_less_2w = this |
|
2112 |
||
2113 |
have PI_diff: "\<And>m n. positive_integral (\<lambda>x. Real (?diff (m + n) x)) = |
|
2114 |
positive_integral (\<lambda>x. Real (2 * w x)) - positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>)" |
|
2115 |
using diff w diff_less_2w |
|
2116 |
by (subst positive_integral_diff[symmetric]) |
|
2117 |
(auto simp: integrable_def intro!: positive_integral_cong) |
|
2118 |
||
2119 |
have "integrable (\<lambda>x. 2 * w x)" |
|
2120 |
using w by (auto intro: integral_cmult) |
|
2121 |
hence I2w_fin: "positive_integral (\<lambda>x. Real (2 * w x)) \<noteq> \<omega>" and |
|
2122 |
borel_2w: "(\<lambda>x. Real (2 * w x)) \<in> borel_measurable M" |
|
2123 |
unfolding integrable_def by auto |
|
2124 |
||
2125 |
have "(INF n. SUP m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>)) = 0" (is "?lim_SUP = 0") |
|
2126 |
proof cases |
|
2127 |
assume eq_0: "positive_integral (\<lambda>x. Real (2 * w x)) = 0" |
|
2128 |
have "\<And>i. positive_integral (\<lambda>x. Real \<bar>u i x - u' x\<bar>) \<le> positive_integral (\<lambda>x. Real (2 * w x))" |
|
2129 |
proof (rule positive_integral_mono) |
|
2130 |
fix i x assume "x \<in> space M" from diff_less_2w[OF this, of i] |
|
2131 |
show "Real \<bar>u i x - u' x\<bar> \<le> Real (2 * w x)" by auto |
|
2132 |
qed |
|
2133 |
hence "\<And>i. positive_integral (\<lambda>x. Real \<bar>u i x - u' x\<bar>) = 0" using eq_0 by auto |
|
2134 |
thus ?thesis by simp |
|
2135 |
next |
|
2136 |
assume neq_0: "positive_integral (\<lambda>x. Real (2 * w x)) \<noteq> 0" |
|
2137 |
have "positive_integral (\<lambda>x. Real (2 * w x)) = positive_integral (SUP n. INF m. (\<lambda>x. Real (?diff (m + n) x)))" |
|
2138 |
proof (rule positive_integral_cong, unfold SUPR_fun_expand INFI_fun_expand, subst add_commute) |
|
2139 |
fix x assume x: "x \<in> space M" |
|
2140 |
show "Real (2 * w x) = (SUP n. INF m. Real (?diff (n + m) x))" |
|
2141 |
proof (rule LIMSEQ_imp_lim_INF[symmetric]) |
|
2142 |
fix j show "0 \<le> ?diff j x" using diff_less_2w[OF x, of j] by simp |
|
2143 |
next |
|
2144 |
have "(\<lambda>i. ?diff i x) ----> 2 * w x - \<bar>u' x - u' x\<bar>" |
|
2145 |
using u'[OF x] by (safe intro!: LIMSEQ_diff LIMSEQ_const LIMSEQ_imp_rabs) |
|
2146 |
thus "(\<lambda>i. ?diff i x) ----> 2 * w x" by simp |
|
2147 |
qed |
|
2148 |
qed |
|
2149 |
also have "\<dots> \<le> (SUP n. INF m. positive_integral (\<lambda>x. Real (?diff (m + n) x)))" |
|
2150 |
using u'_borel w u unfolding integrable_def |
|
2151 |
by (auto intro!: positive_integral_lim_INF) |
|
2152 |
also have "\<dots> = positive_integral (\<lambda>x. Real (2 * w x)) - |
|
2153 |
(INF n. SUP m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>))" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
2154 |
unfolding PI_diff pextreal_INF_minus[OF I2w_fin] pextreal_SUP_minus .. |
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
2155 |
finally show ?thesis using neq_0 I2w_fin by (rule pextreal_le_minus_imp_0) |
38656 | 2156 |
qed |
2157 |
||
2158 |
have [simp]: "\<And>n m x. Real (- \<bar>u (m + n) x - u' x\<bar>) = 0" by auto |
|
2159 |
||
2160 |
have [simp]: "\<And>n m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>) = |
|
2161 |
Real (integral (\<lambda>x. \<bar>u (n + m) x - u' x\<bar>))" |
|
2162 |
using diff by (subst add_commute) (simp add: integral_def integrable_def Real_real) |
|
2163 |
||
2164 |
have "(SUP n. INF m. positive_integral (\<lambda>x. Real \<bar>u (m + n) x - u' x\<bar>)) \<le> ?lim_SUP" |
|
2165 |
(is "?lim_INF \<le> _") by (subst (1 2) add_commute) (rule lim_INF_le_lim_SUP) |
|
2166 |
hence "?lim_INF = Real 0" "?lim_SUP = Real 0" using `?lim_SUP = 0` by auto |
|
2167 |
thus ?lim_diff using diff by (auto intro!: integral_positive lim_INF_eq_lim_SUP) |
|
2168 |
||
2169 |
show ?lim |
|
2170 |
proof (rule LIMSEQ_I) |
|
2171 |
fix r :: real assume "0 < r" |
|
2172 |
from LIMSEQ_D[OF `?lim_diff` this] |
|
2173 |
obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> integral (\<lambda>x. \<bar>u n x - u' x\<bar>) < r" |
|
2174 |
using diff by (auto simp: integral_positive) |
|
2175 |
||
2176 |
show "\<exists>N. \<forall>n\<ge>N. norm (integral (u n) - integral u') < r" |
|
2177 |
proof (safe intro!: exI[of _ N]) |
|
2178 |
fix n assume "N \<le> n" |
|
2179 |
have "\<bar>integral (u n) - integral u'\<bar> = \<bar>integral (\<lambda>x. u n x - u' x)\<bar>" |
|
2180 |
using u `integrable u'` by (auto simp: integral_diff) |
|
2181 |
also have "\<dots> \<le> integral (\<lambda>x. \<bar>u n x - u' x\<bar>)" using u `integrable u'` |
|
2182 |
by (rule_tac integral_triangle_inequality) (auto intro!: integral_diff) |
|
2183 |
also note N[OF `N \<le> n`] |
|
2184 |
finally show "norm (integral (u n) - integral u') < r" by simp |
|
2185 |
qed |
|
2186 |
qed |
|
2187 |
qed |
|
2188 |
||
2189 |
lemma (in measure_space) integral_sums: |
|
2190 |
assumes borel: "\<And>i. integrable (f i)" |
|
2191 |
and summable: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>f i x\<bar>)" |
|
2192 |
and sums: "summable (\<lambda>i. integral (\<lambda>x. \<bar>f i x\<bar>))" |
|
2193 |
shows "integrable (\<lambda>x. (\<Sum>i. f i x))" (is "integrable ?S") |
|
2194 |
and "(\<lambda>i. integral (f i)) sums integral (\<lambda>x. (\<Sum>i. f i x))" (is ?integral) |
|
2195 |
proof - |
|
2196 |
have "\<forall>x\<in>space M. \<exists>w. (\<lambda>i. \<bar>f i x\<bar>) sums w" |
|
2197 |
using summable unfolding summable_def by auto |
|
2198 |
from bchoice[OF this] |
|
2199 |
obtain w where w: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. \<bar>f i x\<bar>) sums w x" by auto |
|
2200 |
||
2201 |
let "?w y" = "if y \<in> space M then w y else 0" |
|
2202 |
||
2203 |
obtain x where abs_sum: "(\<lambda>i. integral (\<lambda>x. \<bar>f i x\<bar>)) sums x" |
|
2204 |
using sums unfolding summable_def .. |
|
2205 |
||
2206 |
have 1: "\<And>n. integrable (\<lambda>x. \<Sum>i = 0..<n. f i x)" |
|
2207 |
using borel by (auto intro!: integral_setsum) |
|
2208 |
||
2209 |
{ fix j x assume [simp]: "x \<in> space M" |
|
2210 |
have "\<bar>\<Sum>i = 0..< j. f i x\<bar> \<le> (\<Sum>i = 0..< j. \<bar>f i x\<bar>)" by (rule setsum_abs) |
|
2211 |
also have "\<dots> \<le> w x" using w[of x] series_pos_le[of "\<lambda>i. \<bar>f i x\<bar>"] unfolding sums_iff by auto |
|
2212 |
finally have "\<bar>\<Sum>i = 0..<j. f i x\<bar> \<le> ?w x" by simp } |
|
2213 |
note 2 = this |
|
2214 |
||
2215 |
have 3: "integrable ?w" |
|
2216 |
proof (rule integral_monotone_convergence(1)) |
|
2217 |
let "?F n y" = "(\<Sum>i = 0..<n. \<bar>f i y\<bar>)" |
|
2218 |
let "?w' n y" = "if y \<in> space M then ?F n y else 0" |
|
2219 |
have "\<And>n. integrable (?F n)" |
|
2220 |
using borel by (auto intro!: integral_setsum integrable_abs) |
|
2221 |
thus "\<And>n. integrable (?w' n)" by (simp cong: integrable_cong) |
|
2222 |
show "mono ?w'" |
|
2223 |
by (auto simp: mono_def le_fun_def intro!: setsum_mono2) |
|
2224 |
{ fix x show "(\<lambda>n. ?w' n x) ----> ?w x" |
|
2225 |
using w by (cases "x \<in> space M") (simp_all add: LIMSEQ_const sums_def) } |
|
2226 |
have *: "\<And>n. integral (?w' n) = (\<Sum>i = 0..< n. integral (\<lambda>x. \<bar>f i x\<bar>))" |
|
2227 |
using borel by (simp add: integral_setsum integrable_abs cong: integral_cong) |
|
2228 |
from abs_sum |
|
2229 |
show "(\<lambda>i. integral (?w' i)) ----> x" unfolding * sums_def . |
|
2230 |
qed |
|
2231 |
||
2232 |
have 4: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> ?w x" using 2[of _ 0] by simp |
|
2233 |
||
2234 |
from summable[THEN summable_rabs_cancel] |
|
2235 |
have 5: "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>n. \<Sum>i = 0..<n. f i x) ----> (\<Sum>i. f i x)" |
|
2236 |
by (auto intro: summable_sumr_LIMSEQ_suminf) |
|
2237 |
||
2238 |
note int = integral_dominated_convergence(1,3)[OF 1 2 3 4 5] |
|
2239 |
||
2240 |
from int show "integrable ?S" by simp |
|
2241 |
||
2242 |
show ?integral unfolding sums_def integral_setsum(1)[symmetric, OF borel] |
|
2243 |
using int(2) by simp |
|
36624 | 2244 |
qed |
2245 |
||
35748 | 2246 |
section "Lebesgue integration on countable spaces" |
2247 |
||
38656 | 2248 |
lemma (in measure_space) integral_on_countable: |
2249 |
assumes f: "f \<in> borel_measurable M" |
|
35748 | 2250 |
and bij: "bij_betw enum S (f ` space M)" |
2251 |
and enum_zero: "enum ` (-S) \<subseteq> {0}" |
|
38656 | 2252 |
and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>" |
2253 |
and abs_summable: "summable (\<lambda>r. \<bar>enum r * real (\<mu> (f -` {enum r} \<inter> space M))\<bar>)" |
|
35748 | 2254 |
shows "integrable f" |
38656 | 2255 |
and "(\<lambda>r. enum r * real (\<mu> (f -` {enum r} \<inter> space M))) sums integral f" (is ?sums) |
35748 | 2256 |
proof - |
38656 | 2257 |
let "?A r" = "f -` {enum r} \<inter> space M" |
2258 |
let "?F r x" = "enum r * indicator (?A r) x" |
|
2259 |
have enum_eq: "\<And>r. enum r * real (\<mu> (?A r)) = integral (?F r)" |
|
2260 |
using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) |
|
35748 | 2261 |
|
38656 | 2262 |
{ fix x assume "x \<in> space M" |
2263 |
hence "f x \<in> enum ` S" using bij unfolding bij_betw_def by auto |
|
2264 |
then obtain i where "i\<in>S" "enum i = f x" by auto |
|
2265 |
have F: "\<And>j. ?F j x = (if j = i then f x else 0)" |
|
2266 |
proof cases |
|
2267 |
fix j assume "j = i" |
|
2268 |
thus "?thesis j" using `x \<in> space M` `enum i = f x` by (simp add: indicator_def) |
|
2269 |
next |
|
2270 |
fix j assume "j \<noteq> i" |
|
2271 |
show "?thesis j" using bij `i \<in> S` `j \<noteq> i` `enum i = f x` enum_zero |
|
2272 |
by (cases "j \<in> S") (auto simp add: indicator_def bij_betw_def inj_on_def) |
|
2273 |
qed |
|
2274 |
hence F_abs: "\<And>j. \<bar>if j = i then f x else 0\<bar> = (if j = i then \<bar>f x\<bar> else 0)" by auto |
|
2275 |
have "(\<lambda>i. ?F i x) sums f x" |
|
2276 |
"(\<lambda>i. \<bar>?F i x\<bar>) sums \<bar>f x\<bar>" |
|
2277 |
by (auto intro!: sums_single simp: F F_abs) } |
|
2278 |
note F_sums_f = this(1) and F_abs_sums_f = this(2) |
|
35748 | 2279 |
|
38656 | 2280 |
have int_f: "integral f = integral (\<lambda>x. \<Sum>r. ?F r x)" "integrable f = integrable (\<lambda>x. \<Sum>r. ?F r x)" |
2281 |
using F_sums_f by (auto intro!: integral_cong integrable_cong simp: sums_iff) |
|
35748 | 2282 |
|
2283 |
{ fix r |
|
38656 | 2284 |
have "integral (\<lambda>x. \<bar>?F r x\<bar>) = integral (\<lambda>x. \<bar>enum r\<bar> * indicator (?A r) x)" |
2285 |
by (auto simp: indicator_def intro!: integral_cong) |
|
2286 |
also have "\<dots> = \<bar>enum r\<bar> * real (\<mu> (?A r))" |
|
2287 |
using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) |
|
2288 |
finally have "integral (\<lambda>x. \<bar>?F r x\<bar>) = \<bar>enum r * real (\<mu> (?A r))\<bar>" |
|
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
2289 |
by (simp add: abs_mult_pos real_pextreal_pos) } |
38656 | 2290 |
note int_abs_F = this |
35748 | 2291 |
|
38656 | 2292 |
have 1: "\<And>i. integrable (\<lambda>x. ?F i x)" |
2293 |
using f fin by (simp add: borel_measurable_vimage integral_cmul_indicator) |
|
2294 |
||
2295 |
have 2: "\<And>x. x \<in> space M \<Longrightarrow> summable (\<lambda>i. \<bar>?F i x\<bar>)" |
|
2296 |
using F_abs_sums_f unfolding sums_iff by auto |
|
2297 |
||
2298 |
from integral_sums(2)[OF 1 2, unfolded int_abs_F, OF _ abs_summable] |
|
2299 |
show ?sums unfolding enum_eq int_f by simp |
|
2300 |
||
2301 |
from integral_sums(1)[OF 1 2, unfolded int_abs_F, OF _ abs_summable] |
|
2302 |
show "integrable f" unfolding int_f by simp |
|
35748 | 2303 |
qed |
2304 |
||
35692 | 2305 |
section "Lebesgue integration on finite space" |
2306 |
||
38656 | 2307 |
lemma (in measure_space) integral_on_finite: |
2308 |
assumes f: "f \<in> borel_measurable M" and finite: "finite (f`space M)" |
|
2309 |
and fin: "\<And>x. x \<noteq> 0 \<Longrightarrow> \<mu> (f -` {x} \<inter> space M) \<noteq> \<omega>" |
|
2310 |
shows "integrable f" |
|
2311 |
and "integral (\<lambda>x. f x) = |
|
2312 |
(\<Sum> r \<in> f`space M. r * real (\<mu> (f -` {r} \<inter> space M)))" (is "?integral") |
|
35582 | 2313 |
proof - |
38656 | 2314 |
let "?A r" = "f -` {r} \<inter> space M" |
2315 |
let "?S x" = "\<Sum>r\<in>f`space M. r * indicator (?A r) x" |
|
35582 | 2316 |
|
38656 | 2317 |
{ fix x assume "x \<in> space M" |
2318 |
have "f x = (\<Sum>r\<in>f`space M. if x \<in> ?A r then r else 0)" |
|
2319 |
using finite `x \<in> space M` by (simp add: setsum_cases) |
|
2320 |
also have "\<dots> = ?S x" |
|
2321 |
by (auto intro!: setsum_cong) |
|
2322 |
finally have "f x = ?S x" . } |
|
2323 |
note f_eq = this |
|
2324 |
||
2325 |
have f_eq_S: "integrable f \<longleftrightarrow> integrable ?S" "integral f = integral ?S" |
|
2326 |
by (auto intro!: integrable_cong integral_cong simp only: f_eq) |
|
2327 |
||
2328 |
show "integrable f" ?integral using fin f f_eq_S |
|
2329 |
by (simp_all add: integral_cmul_indicator borel_measurable_vimage) |
|
35582 | 2330 |
qed |
2331 |
||
39092 | 2332 |
lemma (in finite_measure_space) simple_function_finite[simp, intro]: "simple_function f" |
40875 | 2333 |
unfolding simple_function_def using finite_space by auto |
35977 | 2334 |
|
38705 | 2335 |
lemma (in finite_measure_space) borel_measurable_finite[intro, simp]: "f \<in> borel_measurable M" |
39092 | 2336 |
by (auto intro: borel_measurable_simple_function) |
38705 | 2337 |
|
2338 |
lemma (in finite_measure_space) positive_integral_finite_eq_setsum: |
|
2339 |
"positive_integral f = (\<Sum>x \<in> space M. f x * \<mu> {x})" |
|
2340 |
proof - |
|
2341 |
have *: "positive_integral f = positive_integral (\<lambda>x. \<Sum>y\<in>space M. f y * indicator {y} x)" |
|
2342 |
by (auto intro!: positive_integral_cong simp add: indicator_def if_distrib setsum_cases[OF finite_space]) |
|
2343 |
show ?thesis unfolding * using borel_measurable_finite[of f] |
|
40875 | 2344 |
by (simp add: positive_integral_setsum positive_integral_cmult_indicator) |
38705 | 2345 |
qed |
2346 |
||
35977 | 2347 |
lemma (in finite_measure_space) integral_finite_singleton: |
38656 | 2348 |
shows "integrable f" |
2349 |
and "integral f = (\<Sum>x \<in> space M. f x * real (\<mu> {x}))" (is ?I) |
|
35977 | 2350 |
proof - |
38705 | 2351 |
have [simp]: |
2352 |
"positive_integral (\<lambda>x. Real (f x)) = (\<Sum>x \<in> space M. Real (f x) * \<mu> {x})" |
|
2353 |
"positive_integral (\<lambda>x. Real (- f x)) = (\<Sum>x \<in> space M. Real (- f x) * \<mu> {x})" |
|
2354 |
unfolding positive_integral_finite_eq_setsum by auto |
|
2355 |
show "integrable f" using finite_space finite_measure |
|
40875 | 2356 |
by (simp add: setsum_\<omega> integrable_def) |
38705 | 2357 |
show ?I using finite_measure |
41023
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
2358 |
apply (simp add: integral_def real_of_pextreal_setsum[symmetric] |
9118eb4eb8dc
it is known as the extended reals, not the infinite reals
hoelzl
parents:
40875
diff
changeset
|
2359 |
real_of_pextreal_mult[symmetric] setsum_subtractf[symmetric]) |
38705 | 2360 |
by (rule setsum_cong) (simp_all split: split_if) |
35977 | 2361 |
qed |
2362 |
||
35748 | 2363 |
end |