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