| author | huffman | 
| Tue, 18 May 2010 06:28:42 -0700 | |
| changeset 36978 | 4ec5131c6f46 | 
| parent 36844 | 5f9385ecc1a7 | 
| child 37032 | 58a0757031dd | 
| permissions | -rw-r--r-- | 
| 35582 | 1  | 
header {*Lebesgue Integration*}
 | 
2  | 
||
3  | 
theory Lebesgue  | 
|
4  | 
imports Measure Borel  | 
|
5  | 
begin  | 
|
6  | 
||
7  | 
text{*From the HOL4 Hurd/Coble Lebesgue integration, translated by Armin Heller and Johannes Hoelzl.*}
 | 
|
8  | 
||
9  | 
definition  | 
|
10  | 
"pos_part f = (\<lambda>x. max 0 (f x))"  | 
|
11  | 
||
12  | 
definition  | 
|
13  | 
"neg_part f = (\<lambda>x. - min 0 (f x))"  | 
|
14  | 
||
15  | 
definition  | 
|
16  | 
"nonneg f = (\<forall>x. 0 \<le> f x)"  | 
|
17  | 
||
18  | 
lemma nonneg_pos_part[intro!]:  | 
|
19  | 
  fixes f :: "'c \<Rightarrow> 'd\<Colon>{linorder,zero}"
 | 
|
20  | 
shows "nonneg (pos_part f)"  | 
|
21  | 
unfolding nonneg_def pos_part_def by auto  | 
|
22  | 
||
23  | 
lemma nonneg_neg_part[intro!]:  | 
|
24  | 
  fixes f :: "'c \<Rightarrow> 'd\<Colon>{linorder,ordered_ab_group_add}"
 | 
|
25  | 
shows "nonneg (neg_part f)"  | 
|
26  | 
unfolding nonneg_def neg_part_def min_def by auto  | 
|
27  | 
||
| 36624 | 28  | 
lemma pos_neg_part_abs:  | 
29  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
30  | 
shows "pos_part f x + neg_part f x = \<bar>f x\<bar>"  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
31  | 
unfolding abs_if pos_part_def neg_part_def by auto  | 
| 36624 | 32  | 
|
33  | 
lemma pos_part_abs:  | 
|
34  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
35  | 
shows "pos_part (\<lambda> x. \<bar>f x\<bar>) y = \<bar>f y\<bar>"  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
36  | 
unfolding pos_part_def abs_if by auto  | 
| 36624 | 37  | 
|
38  | 
lemma neg_part_abs:  | 
|
39  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
40  | 
shows "neg_part (\<lambda> x. \<bar>f x\<bar>) y = 0"  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
41  | 
unfolding neg_part_def abs_if by auto  | 
| 36624 | 42  | 
|
| 35692 | 43  | 
lemma (in measure_space)  | 
44  | 
assumes "f \<in> borel_measurable M"  | 
|
45  | 
shows pos_part_borel_measurable: "pos_part f \<in> borel_measurable M"  | 
|
46  | 
and neg_part_borel_measurable: "neg_part f \<in> borel_measurable M"  | 
|
47  | 
using assms  | 
|
48  | 
proof -  | 
|
49  | 
  { fix a :: real
 | 
|
50  | 
    { assume asm: "0 \<le> a"
 | 
|
51  | 
from asm have pp: "\<And> w. (pos_part f w \<le> a) = (f w \<le> a)" unfolding pos_part_def by auto  | 
|
52  | 
      have "{w | w. w \<in> space M \<and> f w \<le> a} \<in> sets M"
 | 
|
53  | 
unfolding pos_part_def using assms borel_measurable_le_iff by auto  | 
|
54  | 
      hence "{w . w \<in> space M \<and> pos_part f w \<le> a} \<in> sets M" using pp by auto }
 | 
|
55  | 
    moreover have "a < 0 \<Longrightarrow> {w \<in> space M. pos_part f w \<le> a} \<in> sets M"
 | 
|
56  | 
unfolding pos_part_def using empty_sets by auto  | 
|
57  | 
    ultimately have "{w . w \<in> space M \<and> pos_part f w \<le> a} \<in> sets M"
 | 
|
58  | 
using le_less_linear by auto  | 
|
59  | 
} hence pos: "pos_part f \<in> borel_measurable M" using borel_measurable_le_iff by auto  | 
|
60  | 
  { fix a :: real
 | 
|
61  | 
    { assume asm: "0 \<le> a"
 | 
|
62  | 
from asm have pp: "\<And> w. (neg_part f w \<le> a) = (f w \<ge> - a)" unfolding neg_part_def by auto  | 
|
63  | 
      have "{w | w. w \<in> space M \<and> f w \<ge> - a} \<in> sets M"
 | 
|
64  | 
unfolding neg_part_def using assms borel_measurable_ge_iff by auto  | 
|
65  | 
      hence "{w . w \<in> space M \<and> neg_part f w \<le> a} \<in> sets M" using pp by auto }
 | 
|
66  | 
    moreover have "a < 0 \<Longrightarrow> {w \<in> space M. neg_part f w \<le> a} = {}" unfolding neg_part_def by auto
 | 
|
67  | 
    moreover hence "a < 0 \<Longrightarrow> {w \<in> space M. neg_part f w \<le> a} \<in> sets M" by (simp only: empty_sets)
 | 
|
68  | 
    ultimately have "{w . w \<in> space M \<and> neg_part f w \<le> a} \<in> sets M"
 | 
|
69  | 
using le_less_linear by auto  | 
|
70  | 
} hence neg: "neg_part f \<in> borel_measurable M" using borel_measurable_le_iff by auto  | 
|
71  | 
from pos neg show "pos_part f \<in> borel_measurable M" and "neg_part f \<in> borel_measurable M" by auto  | 
|
72  | 
qed  | 
|
73  | 
||
74  | 
lemma (in measure_space) pos_part_neg_part_borel_measurable_iff:  | 
|
75  | 
"f \<in> borel_measurable M \<longleftrightarrow>  | 
|
76  | 
pos_part f \<in> borel_measurable M \<and> neg_part f \<in> borel_measurable M"  | 
|
77  | 
proof -  | 
|
78  | 
  { fix x
 | 
|
79  | 
have "f x = pos_part f x - neg_part f x"  | 
|
80  | 
unfolding pos_part_def neg_part_def unfolding max_def min_def  | 
|
81  | 
by auto }  | 
|
82  | 
hence "(\<lambda> x. f x) = (\<lambda> x. pos_part f x - neg_part f x)" by auto  | 
|
83  | 
hence "f = (\<lambda> x. pos_part f x - neg_part f x)" by blast  | 
|
84  | 
from pos_part_borel_measurable[of f] neg_part_borel_measurable[of f]  | 
|
85  | 
borel_measurable_diff_borel_measurable[of "pos_part f" "neg_part f"]  | 
|
86  | 
this  | 
|
87  | 
show ?thesis by auto  | 
|
88  | 
qed  | 
|
89  | 
||
| 35582 | 90  | 
context measure_space  | 
91  | 
begin  | 
|
92  | 
||
| 35692 | 93  | 
section "Simple discrete step function"  | 
94  | 
||
| 35582 | 95  | 
definition  | 
96  | 
"pos_simple f =  | 
|
97  | 
  { (s :: nat set, a, x).
 | 
|
98  | 
finite s \<and> nonneg f \<and> nonneg x \<and> a ` s \<subseteq> sets M \<and>  | 
|
99  | 
(\<forall>t \<in> space M. (\<exists>!i\<in>s. t\<in>a i) \<and> (\<forall>i\<in>s. t \<in> a i \<longrightarrow> f t = x i)) }"  | 
|
100  | 
||
101  | 
definition  | 
|
102  | 
"pos_simple_integral \<equiv> \<lambda>(s, a, x). \<Sum> i \<in> s. x i * measure M (a i)"  | 
|
103  | 
||
104  | 
definition  | 
|
105  | 
"psfis f = pos_simple_integral ` (pos_simple f)"  | 
|
106  | 
||
107  | 
lemma pos_simpleE[consumes 1]:  | 
|
108  | 
assumes ps: "(s, a, x) \<in> pos_simple f"  | 
|
109  | 
obtains "finite s" and "nonneg f" and "nonneg x"  | 
|
110  | 
and "a ` s \<subseteq> sets M" and "(\<Union>i\<in>s. a i) = space M"  | 
|
111  | 
and "disjoint_family_on a s"  | 
|
112  | 
and "\<And>t. t \<in> space M \<Longrightarrow> (\<exists>!i. i \<in> s \<and> t \<in> a i)"  | 
|
113  | 
and "\<And>t i. \<lbrakk> t \<in> space M ; i \<in> s ; t \<in> a i \<rbrakk> \<Longrightarrow> f t = x i"  | 
|
114  | 
proof  | 
|
115  | 
show "finite s" and "nonneg f" and "nonneg x"  | 
|
116  | 
and as_in_M: "a ` s \<subseteq> sets M"  | 
|
117  | 
and *: "\<And>t. t \<in> space M \<Longrightarrow> (\<exists>!i. i \<in> s \<and> t \<in> a i)"  | 
|
118  | 
and **: "\<And>t i. \<lbrakk> t \<in> space M ; i \<in> s ; t \<in> a i \<rbrakk> \<Longrightarrow> f t = x i"  | 
|
119  | 
using ps unfolding pos_simple_def by auto  | 
|
120  | 
||
121  | 
thus t: "(\<Union>i\<in>s. a i) = space M"  | 
|
122  | 
proof safe  | 
|
123  | 
fix x assume "x \<in> space M"  | 
|
124  | 
from *[OF this] show "x \<in> (\<Union>i\<in>s. a i)" by auto  | 
|
125  | 
next  | 
|
126  | 
fix t i assume "i \<in> s"  | 
|
127  | 
hence "a i \<in> sets M" using as_in_M by auto  | 
|
128  | 
moreover assume "t \<in> a i"  | 
|
129  | 
ultimately show "t \<in> space M" using sets_into_space by blast  | 
|
130  | 
qed  | 
|
131  | 
||
132  | 
show "disjoint_family_on a s" unfolding disjoint_family_on_def  | 
|
133  | 
proof safe  | 
|
134  | 
fix i j and t assume "i \<in> s" "t \<in> a i" and "j \<in> s" "t \<in> a j" and "i \<noteq> j"  | 
|
135  | 
    with t * show "t \<in> {}" by auto
 | 
|
136  | 
qed  | 
|
137  | 
qed  | 
|
138  | 
||
139  | 
lemma pos_simple_cong:  | 
|
140  | 
assumes "nonneg f" and "nonneg g" and "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"  | 
|
141  | 
shows "pos_simple f = pos_simple g"  | 
|
142  | 
unfolding pos_simple_def using assms by auto  | 
|
143  | 
||
144  | 
lemma psfis_cong:  | 
|
145  | 
assumes "nonneg f" and "nonneg g" and "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"  | 
|
146  | 
shows "psfis f = psfis g"  | 
|
147  | 
unfolding psfis_def using pos_simple_cong[OF assms] by simp  | 
|
148  | 
||
| 35692 | 149  | 
lemma psfis_0: "0 \<in> psfis (\<lambda>x. 0)"  | 
150  | 
unfolding psfis_def pos_simple_def pos_simple_integral_def  | 
|
151  | 
by (auto simp: nonneg_def  | 
|
152  | 
      intro: image_eqI[where x="({0}, (\<lambda>i. space M), (\<lambda>i. 0))"])
 | 
|
153  | 
||
| 35582 | 154  | 
lemma pos_simple_setsum_indicator_fn:  | 
155  | 
assumes ps: "(s, a, x) \<in> pos_simple f"  | 
|
156  | 
and "t \<in> space M"  | 
|
157  | 
shows "(\<Sum>i\<in>s. x i * indicator_fn (a i) t) = f t"  | 
|
158  | 
proof -  | 
|
159  | 
from assms obtain i where *: "i \<in> s" "t \<in> a i"  | 
|
160  | 
and "finite s" and xi: "x i = f t" by (auto elim!: pos_simpleE)  | 
|
161  | 
||
162  | 
have **: "(\<Sum>i\<in>s. x i * indicator_fn (a i) t) =  | 
|
163  | 
    (\<Sum>j\<in>s. if j \<in> {i} then x i else 0)"
 | 
|
164  | 
unfolding indicator_fn_def using assms *  | 
|
165  | 
by (auto intro!: setsum_cong elim!: pos_simpleE)  | 
|
166  | 
show ?thesis unfolding ** setsum_cases[OF `finite s`] xi  | 
|
167  | 
using `i \<in> s` by simp  | 
|
168  | 
qed  | 
|
169  | 
||
| 35692 | 170  | 
lemma pos_simple_common_partition:  | 
| 35582 | 171  | 
assumes psf: "(s, a, x) \<in> pos_simple f"  | 
172  | 
assumes psg: "(s', b, y) \<in> pos_simple g"  | 
|
173  | 
obtains z z' c k where "(k, c, z) \<in> pos_simple f" "(k, c, z') \<in> pos_simple g"  | 
|
174  | 
proof -  | 
|
175  | 
(* definitions *)  | 
|
176  | 
||
177  | 
  def k == "{0 ..< card (s \<times> s')}"
 | 
|
178  | 
have fs: "finite s" "finite s'" "finite k" unfolding k_def  | 
|
179  | 
using psf psg unfolding pos_simple_def by auto  | 
|
180  | 
hence "finite (s \<times> s')" by simp  | 
|
181  | 
then obtain p where p: "p ` k = s \<times> s'" "inj_on p k"  | 
|
182  | 
using ex_bij_betw_nat_finite[of "s \<times> s'"] unfolding bij_betw_def k_def by blast  | 
|
183  | 
def c == "\<lambda> i. a (fst (p i)) \<inter> b (snd (p i))"  | 
|
184  | 
def z == "\<lambda> i. x (fst (p i))"  | 
|
185  | 
def z' == "\<lambda> i. y (snd (p i))"  | 
|
186  | 
||
187  | 
have "finite k" unfolding k_def by simp  | 
|
188  | 
||
189  | 
have "nonneg z" and "nonneg z'"  | 
|
190  | 
using psf psg unfolding z_def z'_def pos_simple_def nonneg_def by auto  | 
|
191  | 
||
192  | 
have ck_subset_M: "c ` k \<subseteq> sets M"  | 
|
193  | 
proof  | 
|
194  | 
fix x assume "x \<in> c ` k"  | 
|
195  | 
then obtain j where "j \<in> k" and "x = c j" by auto  | 
|
196  | 
hence "p j \<in> s \<times> s'" using p(1) by auto  | 
|
197  | 
hence "a (fst (p j)) \<in> sets M" (is "?a \<in> _")  | 
|
198  | 
and "b (snd (p j)) \<in> sets M" (is "?b \<in> _")  | 
|
199  | 
using psf psg unfolding pos_simple_def by auto  | 
|
200  | 
thus "x \<in> sets M" unfolding `x = c j` c_def using Int by simp  | 
|
201  | 
qed  | 
|
202  | 
||
203  | 
  { fix t assume "t \<in> space M"
 | 
|
204  | 
hence ex1s: "\<exists>!i\<in>s. t \<in> a i" and ex1s': "\<exists>!i\<in>s'. t \<in> b i"  | 
|
205  | 
using psf psg unfolding pos_simple_def by auto  | 
|
206  | 
then obtain j j' where j: "j \<in> s" "t \<in> a j" and j': "j' \<in> s'" "t \<in> b j'"  | 
|
207  | 
by auto  | 
|
208  | 
then obtain i :: nat where i: "(j,j') = p i" and "i \<in> k" using p by auto  | 
|
209  | 
have "\<exists>!i\<in>k. t \<in> c i"  | 
|
210  | 
proof (rule ex1I[of _ i])  | 
|
211  | 
show "\<And>x. x \<in> k \<Longrightarrow> t \<in> c x \<Longrightarrow> x = i"  | 
|
212  | 
proof -  | 
|
213  | 
fix l assume "l \<in> k" "t \<in> c l"  | 
|
214  | 
hence "p l \<in> s \<times> s'" and t_in: "t \<in> a (fst (p l))" "t \<in> b (snd (p l))"  | 
|
215  | 
using p unfolding c_def by auto  | 
|
216  | 
hence "fst (p l) \<in> s" and "snd (p l) \<in> s'" by auto  | 
|
217  | 
with t_in j j' ex1s ex1s' have "p l = (j, j')" by (cases "p l", auto)  | 
|
218  | 
thus "l = i"  | 
|
219  | 
using `(j, j') = p i` p(2)[THEN inj_onD] `l \<in> k` `i \<in> k` by auto  | 
|
220  | 
qed  | 
|
221  | 
||
222  | 
show "i \<in> k \<and> t \<in> c i"  | 
|
223  | 
using `i \<in> k` `t \<in> a j` `t \<in> b j'` c_def i[symmetric] by auto  | 
|
224  | 
qed auto  | 
|
225  | 
} note ex1 = this  | 
|
226  | 
||
227  | 
show thesis  | 
|
228  | 
proof (rule that)  | 
|
229  | 
    { fix t i assume "t \<in> space M" and "i \<in> k"
 | 
|
230  | 
hence "p i \<in> s \<times> s'" using p(1) by auto  | 
|
231  | 
hence "fst (p i) \<in> s" by auto  | 
|
232  | 
moreover  | 
|
233  | 
assume "t \<in> c i"  | 
|
234  | 
hence "t \<in> a (fst (p i))" unfolding c_def by auto  | 
|
235  | 
ultimately have "f t = z i" using psf `t \<in> space M`  | 
|
236  | 
unfolding z_def pos_simple_def by auto  | 
|
237  | 
}  | 
|
238  | 
thus "(k, c, z) \<in> pos_simple f"  | 
|
239  | 
using psf `finite k` `nonneg z` ck_subset_M ex1  | 
|
240  | 
unfolding pos_simple_def by auto  | 
|
241  | 
||
242  | 
    { fix t i assume "t \<in> space M" and "i \<in> k"
 | 
|
243  | 
hence "p i \<in> s \<times> s'" using p(1) by auto  | 
|
244  | 
hence "snd (p i) \<in> s'" by auto  | 
|
245  | 
moreover  | 
|
246  | 
assume "t \<in> c i"  | 
|
247  | 
hence "t \<in> b (snd (p i))" unfolding c_def by auto  | 
|
248  | 
ultimately have "g t = z' i" using psg `t \<in> space M`  | 
|
249  | 
unfolding z'_def pos_simple_def by auto  | 
|
250  | 
}  | 
|
251  | 
thus "(k, c, z') \<in> pos_simple g"  | 
|
252  | 
using psg `finite k` `nonneg z'` ck_subset_M ex1  | 
|
253  | 
unfolding pos_simple_def by auto  | 
|
254  | 
qed  | 
|
255  | 
qed  | 
|
256  | 
||
| 35692 | 257  | 
lemma pos_simple_integral_equal:  | 
| 35582 | 258  | 
assumes psx: "(s, a, x) \<in> pos_simple f"  | 
259  | 
assumes psy: "(s', b, y) \<in> pos_simple f"  | 
|
260  | 
shows "pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)"  | 
|
261  | 
unfolding pos_simple_integral_def  | 
|
262  | 
proof simp  | 
|
263  | 
have "(\<Sum>i\<in>s. x i * measure M (a i)) =  | 
|
264  | 
(\<Sum>i\<in>s. (\<Sum>j \<in> s'. x i * measure M (a i \<inter> b j)))" (is "?left = _")  | 
|
265  | 
using psy psx unfolding setsum_right_distrib[symmetric]  | 
|
266  | 
by (auto intro!: setsum_cong measure_setsum_split elim!: pos_simpleE)  | 
|
267  | 
also have "... = (\<Sum>i\<in>s. (\<Sum>j \<in> s'. y j * measure M (a i \<inter> b j)))"  | 
|
268  | 
proof (rule setsum_cong, simp, rule setsum_cong, simp_all)  | 
|
269  | 
fix i j assume i: "i \<in> s" and j: "j \<in> s'"  | 
|
270  | 
hence "a i \<in> sets M" using psx by (auto elim!: pos_simpleE)  | 
|
271  | 
||
272  | 
show "measure M (a i \<inter> b j) = 0 \<or> x i = y j"  | 
|
273  | 
    proof (cases "a i \<inter> b j = {}")
 | 
|
274  | 
case True thus ?thesis using empty_measure by simp  | 
|
275  | 
next  | 
|
276  | 
case False then obtain t where t: "t \<in> a i" "t \<in> b j" by auto  | 
|
277  | 
hence "t \<in> space M" using `a i \<in> sets M` sets_into_space by auto  | 
|
278  | 
with psx psy t i j have "x i = f t" and "y j = f t"  | 
|
279  | 
unfolding pos_simple_def by auto  | 
|
280  | 
thus ?thesis by simp  | 
|
281  | 
qed  | 
|
282  | 
qed  | 
|
283  | 
also have "... = (\<Sum>j\<in>s'. (\<Sum>i\<in>s. y j * measure M (a i \<inter> b j)))"  | 
|
284  | 
by (subst setsum_commute) simp  | 
|
285  | 
also have "... = (\<Sum>i\<in>s'. y i * measure M (b i))" (is "?sum_sum = ?right")  | 
|
286  | 
proof (rule setsum_cong)  | 
|
287  | 
fix j assume "j \<in> s'"  | 
|
288  | 
have "y j * measure M (b j) = (\<Sum>i\<in>s. y j * measure M (b j \<inter> a i))"  | 
|
289  | 
using psx psy `j \<in> s'` unfolding setsum_right_distrib[symmetric]  | 
|
290  | 
by (auto intro!: measure_setsum_split elim!: pos_simpleE)  | 
|
291  | 
thus "(\<Sum>i\<in>s. y j * measure M (a i \<inter> b j)) = y j * measure M (b j)"  | 
|
292  | 
by (auto intro!: setsum_cong arg_cong[where f="measure M"])  | 
|
293  | 
qed simp  | 
|
294  | 
finally show "?left = ?right" .  | 
|
295  | 
qed  | 
|
296  | 
||
| 35692 | 297  | 
lemma psfis_present:  | 
| 35582 | 298  | 
assumes "A \<in> psfis f"  | 
299  | 
assumes "B \<in> psfis g"  | 
|
300  | 
obtains z z' c k where  | 
|
301  | 
"A = pos_simple_integral (k, c, z)"  | 
|
302  | 
"B = pos_simple_integral (k, c, z')"  | 
|
303  | 
"(k, c, z) \<in> pos_simple f"  | 
|
304  | 
"(k, c, z') \<in> pos_simple g"  | 
|
305  | 
using assms  | 
|
306  | 
proof -  | 
|
307  | 
from assms obtain s a x s' b y where  | 
|
308  | 
ps: "(s, a, x) \<in> pos_simple f" "(s', b, y) \<in> pos_simple g" and  | 
|
309  | 
A: "A = pos_simple_integral (s, a, x)" and  | 
|
310  | 
B: "B = pos_simple_integral (s', b, y)"  | 
|
311  | 
unfolding psfis_def pos_simple_integral_def by auto  | 
|
312  | 
||
313  | 
guess z z' c k using pos_simple_common_partition[OF ps] . note part = this  | 
|
314  | 
show thesis  | 
|
315  | 
proof (rule that[of k c z z', OF _ _ part])  | 
|
316  | 
show "A = pos_simple_integral (k, c, z)"  | 
|
317  | 
unfolding A by (rule pos_simple_integral_equal[OF ps(1) part(1)])  | 
|
318  | 
show "B = pos_simple_integral (k, c, z')"  | 
|
319  | 
unfolding B by (rule pos_simple_integral_equal[OF ps(2) part(2)])  | 
|
320  | 
qed  | 
|
321  | 
qed  | 
|
322  | 
||
| 35692 | 323  | 
lemma pos_simple_integral_add:  | 
| 35582 | 324  | 
assumes "(s, a, x) \<in> pos_simple f"  | 
325  | 
assumes "(s', b, y) \<in> pos_simple g"  | 
|
326  | 
obtains s'' c z where  | 
|
327  | 
"(s'', c, z) \<in> pos_simple (\<lambda>x. f x + g x)"  | 
|
328  | 
"(pos_simple_integral (s, a, x) +  | 
|
329  | 
pos_simple_integral (s', b, y) =  | 
|
330  | 
pos_simple_integral (s'', c, z))"  | 
|
331  | 
using assms  | 
|
332  | 
proof -  | 
|
333  | 
guess z z' c k  | 
|
334  | 
by (rule pos_simple_common_partition[OF `(s, a, x) \<in> pos_simple f ` `(s', b, y) \<in> pos_simple g`])  | 
|
335  | 
note kczz' = this  | 
|
336  | 
have x: "pos_simple_integral (s, a, x) = pos_simple_integral (k, c, z)"  | 
|
337  | 
by (rule pos_simple_integral_equal) (fact, fact)  | 
|
338  | 
have y: "pos_simple_integral (s', b, y) = pos_simple_integral (k, c, z')"  | 
|
339  | 
by (rule pos_simple_integral_equal) (fact, fact)  | 
|
340  | 
||
341  | 
have "pos_simple_integral (k, c, (\<lambda> x. z x + z' x))  | 
|
342  | 
= (\<Sum> x \<in> k. (z x + z' x) * measure M (c x))"  | 
|
343  | 
unfolding pos_simple_integral_def by auto  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
344  | 
also have "\<dots> = (\<Sum> x \<in> k. z x * measure M (c x) + z' x * measure M (c x))"  | 
| 
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
345  | 
by (simp add: left_distrib)  | 
| 35582 | 346  | 
also have "\<dots> = (\<Sum> x \<in> k. z x * measure M (c x)) + (\<Sum> x \<in> k. z' x * measure M (c x))" using setsum_addf by auto  | 
347  | 
also have "\<dots> = pos_simple_integral (k, c, z) + pos_simple_integral (k, c, z')" unfolding pos_simple_integral_def by auto  | 
|
348  | 
finally have ths: "pos_simple_integral (s, a, x) + pos_simple_integral (s', b, y) =  | 
|
349  | 
pos_simple_integral (k, c, (\<lambda> x. z x + z' x))" using x y by auto  | 
|
350  | 
show ?thesis  | 
|
351  | 
apply (rule that[of k c "(\<lambda> x. z x + z' x)", OF _ ths])  | 
|
352  | 
using kczz' unfolding pos_simple_def nonneg_def by (auto intro!: add_nonneg_nonneg)  | 
|
353  | 
qed  | 
|
354  | 
||
355  | 
lemma psfis_add:  | 
|
356  | 
assumes "a \<in> psfis f" "b \<in> psfis g"  | 
|
357  | 
shows "a + b \<in> psfis (\<lambda>x. f x + g x)"  | 
|
358  | 
using assms pos_simple_integral_add  | 
|
359  | 
unfolding psfis_def by auto blast  | 
|
360  | 
||
361  | 
lemma pos_simple_integral_mono_on_mspace:  | 
|
362  | 
assumes f: "(s, a, x) \<in> pos_simple f"  | 
|
363  | 
assumes g: "(s', b, y) \<in> pos_simple g"  | 
|
364  | 
assumes mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"  | 
|
365  | 
shows "pos_simple_integral (s, a, x) \<le> pos_simple_integral (s', b, y)"  | 
|
366  | 
using assms  | 
|
367  | 
proof -  | 
|
368  | 
guess z z' c k by (rule pos_simple_common_partition[OF f g])  | 
|
369  | 
note kczz' = this  | 
|
370  | 
  (* w = z and w' = z'  except where c _ = {} or undef *)
 | 
|
371  | 
  def w == "\<lambda> i. if i \<notin> k \<or> c i = {} then 0 else z i"
 | 
|
372  | 
  def w' == "\<lambda> i. if i \<notin> k \<or> c i = {} then 0 else z' i"
 | 
|
373  | 
  { fix i
 | 
|
374  | 
have "w i \<le> w' i"  | 
|
375  | 
    proof (cases "i \<notin> k \<or> c i = {}")
 | 
|
376  | 
      case False hence "i \<in> k" "c i \<noteq> {}" by auto
 | 
|
377  | 
then obtain v where v: "v \<in> c i" and "c i \<in> sets M"  | 
|
378  | 
using kczz'(1) unfolding pos_simple_def by blast  | 
|
379  | 
hence "v \<in> space M" using sets_into_space by blast  | 
|
380  | 
with mono[OF `v \<in> space M`] kczz' `i \<in> k` `v \<in> c i`  | 
|
381  | 
have "z i \<le> z' i" unfolding pos_simple_def by simp  | 
|
382  | 
thus ?thesis unfolding w_def w'_def using False by auto  | 
|
383  | 
next  | 
|
384  | 
case True thus ?thesis unfolding w_def w'_def by auto  | 
|
385  | 
qed  | 
|
386  | 
} note w_mn = this  | 
|
387  | 
||
388  | 
(* some technical stuff for the calculation*)  | 
|
389  | 
have "\<And> i. i \<in> k \<Longrightarrow> c i \<in> sets M" using kczz' unfolding pos_simple_def by auto  | 
|
390  | 
hence "\<And> i. i \<in> k \<Longrightarrow> measure M (c i) \<ge> 0" using positive by auto  | 
|
391  | 
hence w_mono: "\<And> i. i \<in> k \<Longrightarrow> w i * measure M (c i) \<le> w' i * measure M (c i)"  | 
|
392  | 
using mult_right_mono w_mn by blast  | 
|
393  | 
||
394  | 
  { fix i have "\<lbrakk>i \<in> k ; z i \<noteq> w i\<rbrakk> \<Longrightarrow> measure M (c i) = 0"
 | 
|
395  | 
      unfolding w_def by (cases "c i = {}") auto }
 | 
|
396  | 
hence zw: "\<And> i. i \<in> k \<Longrightarrow> z i * measure M (c i) = w i * measure M (c i)" by auto  | 
|
397  | 
||
398  | 
  { fix i have "i \<in> k \<Longrightarrow> z i * measure M (c i) = w i * measure M (c i)"
 | 
|
399  | 
      unfolding w_def by (cases "c i = {}") simp_all }
 | 
|
400  | 
note zw = this  | 
|
401  | 
||
402  | 
  { fix i have "i \<in> k \<Longrightarrow> z' i * measure M (c i) = w' i * measure M (c i)"
 | 
|
403  | 
      unfolding w'_def by (cases "c i = {}") simp_all }
 | 
|
404  | 
note z'w' = this  | 
|
405  | 
||
406  | 
(* the calculation *)  | 
|
407  | 
have "pos_simple_integral (s, a, x) = pos_simple_integral (k, c, z)"  | 
|
408  | 
using f kczz'(1) by (rule pos_simple_integral_equal)  | 
|
409  | 
also have "\<dots> = (\<Sum> i \<in> k. z i * measure M (c i))"  | 
|
410  | 
unfolding pos_simple_integral_def by auto  | 
|
411  | 
also have "\<dots> = (\<Sum> i \<in> k. w i * measure M (c i))"  | 
|
412  | 
using setsum_cong2[of k, OF zw] by auto  | 
|
413  | 
also have "\<dots> \<le> (\<Sum> i \<in> k. w' i * measure M (c i))"  | 
|
414  | 
using setsum_mono[OF w_mono, of k] by auto  | 
|
415  | 
also have "\<dots> = (\<Sum> i \<in> k. z' i * measure M (c i))"  | 
|
416  | 
using setsum_cong2[of k, OF z'w'] by auto  | 
|
417  | 
also have "\<dots> = pos_simple_integral (k, c, z')"  | 
|
418  | 
unfolding pos_simple_integral_def by auto  | 
|
419  | 
also have "\<dots> = pos_simple_integral (s', b, y)"  | 
|
420  | 
using kczz'(2) g by (rule pos_simple_integral_equal)  | 
|
421  | 
finally show "pos_simple_integral (s, a, x) \<le> pos_simple_integral (s', b, y)"  | 
|
422  | 
by simp  | 
|
423  | 
qed  | 
|
424  | 
||
425  | 
lemma pos_simple_integral_mono:  | 
|
426  | 
assumes a: "(s, a, x) \<in> pos_simple f" "(s', b, y) \<in> pos_simple g"  | 
|
427  | 
assumes "\<And> z. f z \<le> g z"  | 
|
428  | 
shows "pos_simple_integral (s, a, x) \<le> pos_simple_integral (s', b, y)"  | 
|
429  | 
using assms pos_simple_integral_mono_on_mspace by auto  | 
|
430  | 
||
431  | 
lemma psfis_mono:  | 
|
432  | 
assumes "a \<in> psfis f" "b \<in> psfis g"  | 
|
433  | 
assumes "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x"  | 
|
434  | 
shows "a \<le> b"  | 
|
435  | 
using assms pos_simple_integral_mono_on_mspace unfolding psfis_def by auto  | 
|
436  | 
||
437  | 
lemma pos_simple_fn_integral_unique:  | 
|
438  | 
assumes "(s, a, x) \<in> pos_simple f" "(s', b, y) \<in> pos_simple f"  | 
|
439  | 
shows "pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)"  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
440  | 
using assms by (rule pos_simple_integral_equal) (* FIXME: redundant lemma *)  | 
| 35582 | 441  | 
|
442  | 
lemma psfis_unique:  | 
|
443  | 
assumes "a \<in> psfis f" "b \<in> psfis f"  | 
|
444  | 
shows "a = b"  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
445  | 
using assms by (intro order_antisym psfis_mono [OF _ _ order_refl])  | 
| 35582 | 446  | 
|
447  | 
lemma pos_simple_integral_indicator:  | 
|
448  | 
assumes "A \<in> sets M"  | 
|
449  | 
obtains s a x where  | 
|
450  | 
"(s, a, x) \<in> pos_simple (indicator_fn A)"  | 
|
451  | 
"measure M A = pos_simple_integral (s, a, x)"  | 
|
452  | 
using assms  | 
|
453  | 
proof -  | 
|
454  | 
  def s == "{0, 1} :: nat set"
 | 
|
455  | 
def a == "\<lambda> i :: nat. if i = 0 then A else space M - A"  | 
|
456  | 
def x == "\<lambda> i :: nat. if i = 0 then 1 else (0 :: real)"  | 
|
457  | 
have eq: "pos_simple_integral (s, a, x) = measure M A"  | 
|
458  | 
unfolding s_def a_def x_def pos_simple_integral_def by auto  | 
|
459  | 
have "(s, a, x) \<in> pos_simple (indicator_fn A)"  | 
|
460  | 
unfolding pos_simple_def indicator_fn_def s_def a_def x_def nonneg_def  | 
|
461  | 
using assms sets_into_space by auto  | 
|
462  | 
from that[OF this] eq show thesis by auto  | 
|
463  | 
qed  | 
|
464  | 
||
465  | 
lemma psfis_indicator:  | 
|
466  | 
assumes "A \<in> sets M"  | 
|
467  | 
shows "measure M A \<in> psfis (indicator_fn A)"  | 
|
468  | 
using pos_simple_integral_indicator[OF assms]  | 
|
469  | 
unfolding psfis_def image_def by auto  | 
|
470  | 
||
471  | 
lemma pos_simple_integral_mult:  | 
|
472  | 
assumes f: "(s, a, x) \<in> pos_simple f"  | 
|
473  | 
assumes "0 \<le> z"  | 
|
474  | 
obtains s' b y where  | 
|
475  | 
"(s', b, y) \<in> pos_simple (\<lambda>x. z * f x)"  | 
|
476  | 
"pos_simple_integral (s', b, y) = z * pos_simple_integral (s, a, x)"  | 
|
477  | 
using assms that[of s a "\<lambda>i. z * x i"]  | 
|
478  | 
by (simp add: pos_simple_def pos_simple_integral_def setsum_right_distrib algebra_simps nonneg_def mult_nonneg_nonneg)  | 
|
479  | 
||
480  | 
lemma psfis_mult:  | 
|
481  | 
assumes "r \<in> psfis f"  | 
|
482  | 
assumes "0 \<le> z"  | 
|
483  | 
shows "z * r \<in> psfis (\<lambda>x. z * f x)"  | 
|
484  | 
using assms  | 
|
485  | 
proof -  | 
|
486  | 
from assms obtain s a x  | 
|
487  | 
where sax: "(s, a, x) \<in> pos_simple f"  | 
|
488  | 
and r: "r = pos_simple_integral (s, a, x)"  | 
|
489  | 
unfolding psfis_def image_def by auto  | 
|
490  | 
obtain s' b y where  | 
|
491  | 
"(s', b, y) \<in> pos_simple (\<lambda>x. z * f x)"  | 
|
492  | 
"z * pos_simple_integral (s, a, x) = pos_simple_integral (s', b, y)"  | 
|
493  | 
using pos_simple_integral_mult[OF sax `0 \<le> z`, of thesis] by auto  | 
|
494  | 
thus ?thesis using r unfolding psfis_def image_def by auto  | 
|
495  | 
qed  | 
|
496  | 
||
497  | 
lemma psfis_setsum_image:  | 
|
498  | 
assumes "finite P"  | 
|
499  | 
assumes "\<And>i. i \<in> P \<Longrightarrow> a i \<in> psfis (f i)"  | 
|
500  | 
shows "(setsum a P) \<in> psfis (\<lambda>t. \<Sum>i \<in> P. f i t)"  | 
|
501  | 
using assms  | 
|
502  | 
proof (induct P)  | 
|
503  | 
case empty  | 
|
504  | 
  let ?s = "{0 :: nat}"
 | 
|
505  | 
  let ?a = "\<lambda> i. if i = (0 :: nat) then space M else {}"
 | 
|
506  | 
let ?x = "\<lambda> (i :: nat). (0 :: real)"  | 
|
507  | 
have "(?s, ?a, ?x) \<in> pos_simple (\<lambda> t. (0 :: real))"  | 
|
508  | 
unfolding pos_simple_def image_def nonneg_def by auto  | 
|
509  | 
moreover have "(\<Sum> i \<in> ?s. ?x i * measure M (?a i)) = 0" by auto  | 
|
510  | 
ultimately have "0 \<in> psfis (\<lambda> t. 0)"  | 
|
511  | 
unfolding psfis_def image_def pos_simple_integral_def nonneg_def  | 
|
512  | 
by (auto intro!:bexI[of _ "(?s, ?a, ?x)"])  | 
|
513  | 
thus ?case by auto  | 
|
514  | 
next  | 
|
515  | 
case (insert x P) note asms = this  | 
|
516  | 
have "finite P" by fact  | 
|
517  | 
have "x \<notin> P" by fact  | 
|
518  | 
have "(\<And>i. i \<in> P \<Longrightarrow> a i \<in> psfis (f i)) \<Longrightarrow>  | 
|
519  | 
setsum a P \<in> psfis (\<lambda>t. \<Sum>i\<in>P. f i t)" by fact  | 
|
520  | 
have "setsum a (insert x P) = a x + setsum a P" using `finite P` `x \<notin> P` by auto  | 
|
521  | 
also have "\<dots> \<in> psfis (\<lambda> t. f x t + (\<Sum> i \<in> P. f i t))"  | 
|
522  | 
using asms psfis_add by auto  | 
|
523  | 
also have "\<dots> = psfis (\<lambda> t. \<Sum> i \<in> insert x P. f i t)"  | 
|
524  | 
using `x \<notin> P` `finite P` by auto  | 
|
525  | 
finally show ?case by simp  | 
|
526  | 
qed  | 
|
527  | 
||
528  | 
lemma psfis_intro:  | 
|
529  | 
assumes "a ` P \<subseteq> sets M" and "nonneg x" and "finite P"  | 
|
530  | 
shows "(\<Sum>i\<in>P. x i * measure M (a i)) \<in> psfis (\<lambda>t. \<Sum>i\<in>P. x i * indicator_fn (a i) t)"  | 
|
531  | 
using assms psfis_mult psfis_indicator  | 
|
532  | 
unfolding image_def nonneg_def  | 
|
533  | 
by (auto intro!:psfis_setsum_image)  | 
|
534  | 
||
535  | 
lemma psfis_nonneg: "a \<in> psfis f \<Longrightarrow> nonneg f"  | 
|
536  | 
unfolding psfis_def pos_simple_def by auto  | 
|
537  | 
||
538  | 
lemma pos_psfis: "r \<in> psfis f \<Longrightarrow> 0 \<le> r"  | 
|
539  | 
unfolding psfis_def pos_simple_integral_def image_def pos_simple_def nonneg_def  | 
|
540  | 
using positive[unfolded positive_def] by (auto intro!:setsum_nonneg mult_nonneg_nonneg)  | 
|
541  | 
||
542  | 
lemma psfis_borel_measurable:  | 
|
543  | 
assumes "a \<in> psfis f"  | 
|
544  | 
shows "f \<in> borel_measurable M"  | 
|
545  | 
using assms  | 
|
546  | 
proof -  | 
|
547  | 
from assms obtain s a' x where sa'x: "(s, a', x) \<in> pos_simple f" and sa'xa: "pos_simple_integral (s, a', x) = a"  | 
|
548  | 
and fs: "finite s"  | 
|
549  | 
unfolding psfis_def pos_simple_integral_def image_def  | 
|
550  | 
by (auto elim:pos_simpleE)  | 
|
551  | 
  { fix w assume "w \<in> space M"
 | 
|
552  | 
hence "(f w \<le> a) = ((\<Sum> i \<in> s. x i * indicator_fn (a' i) w) \<le> a)"  | 
|
553  | 
using pos_simple_setsum_indicator_fn[OF sa'x, of w] by simp  | 
|
554  | 
} hence "\<And> w. (w \<in> space M \<and> f w \<le> a) = (w \<in> space M \<and> (\<Sum> i \<in> s. x i * indicator_fn (a' i) w) \<le> a)"  | 
|
555  | 
by auto  | 
|
556  | 
  { fix i assume "i \<in> s"
 | 
|
557  | 
hence "indicator_fn (a' i) \<in> borel_measurable M"  | 
|
558  | 
using borel_measurable_indicator using sa'x[unfolded pos_simple_def] by auto  | 
|
559  | 
hence "(\<lambda> w. x i * indicator_fn (a' i) w) \<in> borel_measurable M"  | 
|
560  | 
using affine_borel_measurable[of "\<lambda> w. indicator_fn (a' i) w" 0 "x i"]  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
561  | 
by (simp add: mult_commute) }  | 
| 35582 | 562  | 
from borel_measurable_setsum_borel_measurable[OF fs this] affine_borel_measurable  | 
563  | 
have "(\<lambda> w. (\<Sum> i \<in> s. x i * indicator_fn (a' i) w)) \<in> borel_measurable M" by auto  | 
|
564  | 
from borel_measurable_cong[OF pos_simple_setsum_indicator_fn[OF sa'x]] this  | 
|
565  | 
show ?thesis by simp  | 
|
566  | 
qed  | 
|
567  | 
||
568  | 
lemma psfis_mono_conv_mono:  | 
|
569  | 
assumes mono: "mono_convergent u f (space M)"  | 
|
570  | 
and ps_u: "\<And>n. x n \<in> psfis (u n)"  | 
|
571  | 
and "x ----> y"  | 
|
572  | 
and "r \<in> psfis s"  | 
|
573  | 
and f_upper_bound: "\<And>t. t \<in> space M \<Longrightarrow> s t \<le> f t"  | 
|
574  | 
shows "r <= y"  | 
|
575  | 
proof (rule field_le_mult_one_interval)  | 
|
576  | 
fix z :: real assume "0 < z" and "z < 1"  | 
|
577  | 
hence "0 \<le> z" by auto  | 
|
578  | 
  let "?B' n" = "{w \<in> space M. z * s w <= u n w}"
 | 
|
579  | 
||
580  | 
have "incseq x" unfolding incseq_def  | 
|
581  | 
proof safe  | 
|
582  | 
fix m n :: nat assume "m \<le> n"  | 
|
583  | 
show "x m \<le> x n"  | 
|
584  | 
proof (rule psfis_mono[OF `x m \<in> psfis (u m)` `x n \<in> psfis (u n)`])  | 
|
585  | 
fix t assume "t \<in> space M"  | 
|
586  | 
with mono_convergentD[OF mono this] `m \<le> n` show "u m t \<le> u n t"  | 
|
587  | 
unfolding incseq_def by auto  | 
|
588  | 
qed  | 
|
589  | 
qed  | 
|
590  | 
||
591  | 
from `r \<in> psfis s`  | 
|
592  | 
obtain s' a x' where r: "r = pos_simple_integral (s', a, x')"  | 
|
593  | 
and ps_s: "(s', a, x') \<in> pos_simple s"  | 
|
594  | 
unfolding psfis_def by auto  | 
|
595  | 
||
596  | 
  { fix t i assume "i \<in> s'" "t \<in> a i"
 | 
|
597  | 
hence "t \<in> space M" using ps_s by (auto elim!: pos_simpleE) }  | 
|
598  | 
note t_in_space = this  | 
|
599  | 
||
600  | 
  { fix n
 | 
|
601  | 
from psfis_borel_measurable[OF `r \<in> psfis s`]  | 
|
602  | 
psfis_borel_measurable[OF ps_u]  | 
|
603  | 
have "?B' n \<in> sets M"  | 
|
604  | 
by (auto intro!:  | 
|
605  | 
borel_measurable_leq_borel_measurable  | 
|
606  | 
borel_measurable_times_borel_measurable  | 
|
607  | 
borel_measurable_const) }  | 
|
608  | 
note B'_in_M = this  | 
|
609  | 
||
610  | 
  { fix n have "(\<lambda>i. a i \<inter> ?B' n) ` s' \<subseteq> sets M" using B'_in_M ps_s
 | 
|
611  | 
by (auto intro!: Int elim!: pos_simpleE) }  | 
|
612  | 
note B'_inter_a_in_M = this  | 
|
613  | 
||
614  | 
let "?sum n" = "(\<Sum>i\<in>s'. x' i * measure M (a i \<inter> ?B' n))"  | 
|
615  | 
  { fix n
 | 
|
616  | 
have "z * ?sum n \<le> x n"  | 
|
617  | 
proof (rule psfis_mono[OF _ ps_u])  | 
|
618  | 
have *: "\<And>i t. indicator_fn (?B' n) t * (x' i * indicator_fn (a i) t) =  | 
|
619  | 
x' i * indicator_fn (a i \<inter> ?B' n) t" unfolding indicator_fn_def by auto  | 
|
620  | 
have ps': "?sum n \<in> psfis (\<lambda>t. indicator_fn (?B' n) t * (\<Sum>i\<in>s'. x' i * indicator_fn (a i) t))"  | 
|
621  | 
unfolding setsum_right_distrib * using B'_in_M ps_s  | 
|
622  | 
by (auto intro!: psfis_intro Int elim!: pos_simpleE)  | 
|
623  | 
also have "... = psfis (\<lambda>t. indicator_fn (?B' n) t * s t)" (is "psfis ?l = psfis ?r")  | 
|
624  | 
proof (rule psfis_cong)  | 
|
625  | 
show "nonneg ?l" using psfis_nonneg[OF ps'] .  | 
|
626  | 
show "nonneg ?r" using psfis_nonneg[OF `r \<in> psfis s`] unfolding nonneg_def indicator_fn_def by auto  | 
|
627  | 
fix t assume "t \<in> space M"  | 
|
628  | 
show "?l t = ?r t" unfolding pos_simple_setsum_indicator_fn[OF ps_s `t \<in> space M`] ..  | 
|
629  | 
qed  | 
|
630  | 
finally show "z * ?sum n \<in> psfis (\<lambda>t. z * ?r t)" using psfis_mult[OF _ `0 \<le> z`] by simp  | 
|
631  | 
next  | 
|
632  | 
fix t assume "t \<in> space M"  | 
|
633  | 
show "z * (indicator_fn (?B' n) t * s t) \<le> u n t"  | 
|
634  | 
using psfis_nonneg[OF ps_u] unfolding nonneg_def indicator_fn_def by auto  | 
|
635  | 
qed }  | 
|
636  | 
hence *: "\<exists>N. \<forall>n\<ge>N. z * ?sum n \<le> x n" by (auto intro!: exI[of _ 0])  | 
|
637  | 
||
638  | 
show "z * r \<le> y" unfolding r pos_simple_integral_def  | 
|
639  | 
proof (rule LIMSEQ_le[OF mult_right.LIMSEQ `x ----> y` *],  | 
|
640  | 
simp, rule LIMSEQ_setsum, rule mult_right.LIMSEQ)  | 
|
641  | 
fix i assume "i \<in> s'"  | 
|
642  | 
from psfis_nonneg[OF `r \<in> psfis s`, unfolded nonneg_def]  | 
|
643  | 
have "\<And>t. 0 \<le> s t" by simp  | 
|
644  | 
||
645  | 
have *: "(\<Union>j. a i \<inter> ?B' j) = a i"  | 
|
646  | 
proof (safe, simp, safe)  | 
|
647  | 
fix t assume "t \<in> a i"  | 
|
648  | 
thus "t \<in> space M" using t_in_space[OF `i \<in> s'`] by auto  | 
|
649  | 
show "\<exists>j. z * s t \<le> u j t"  | 
|
650  | 
proof (cases "s t = 0")  | 
|
651  | 
case True thus ?thesis  | 
|
652  | 
using psfis_nonneg[OF ps_u] unfolding nonneg_def by auto  | 
|
653  | 
next  | 
|
654  | 
case False with `0 \<le> s t`  | 
|
655  | 
have "0 < s t" by auto  | 
|
656  | 
hence "z * s t < 1 * s t" using `0 < z` `z < 1`  | 
|
657  | 
by (auto intro!: mult_strict_right_mono simp del: mult_1_left)  | 
|
658  | 
also have "... \<le> f t" using f_upper_bound `t \<in> space M` by auto  | 
|
659  | 
finally obtain b where "\<And>j. b \<le> j \<Longrightarrow> z * s t < u j t" using `t \<in> space M`  | 
|
660  | 
using mono_conv_outgrow[of "\<lambda>n. u n t" "f t" "z * s t"]  | 
|
661  | 
using mono_convergentD[OF mono] by auto  | 
|
662  | 
from this[of b] show ?thesis by (auto intro!: exI[of _ b])  | 
|
663  | 
qed  | 
|
664  | 
qed  | 
|
665  | 
||
666  | 
show "(\<lambda>n. measure M (a i \<inter> ?B' n)) ----> measure M (a i)"  | 
|
667  | 
proof (safe intro!:  | 
|
668  | 
monotone_convergence[of "\<lambda>n. a i \<inter> ?B' n", unfolded comp_def *])  | 
|
669  | 
fix n show "a i \<inter> ?B' n \<in> sets M"  | 
|
670  | 
using B'_inter_a_in_M[of n] `i \<in> s'` by auto  | 
|
671  | 
next  | 
|
672  | 
fix j t assume "t \<in> space M" and "z * s t \<le> u j t"  | 
|
673  | 
thus "z * s t \<le> u (Suc j) t"  | 
|
674  | 
using mono_convergentD(1)[OF mono, unfolded incseq_def,  | 
|
675  | 
rule_format, of t j "Suc j"]  | 
|
676  | 
by auto  | 
|
677  | 
qed  | 
|
678  | 
qed  | 
|
679  | 
qed  | 
|
680  | 
||
| 35692 | 681  | 
section "Continuous posititve integration"  | 
682  | 
||
683  | 
definition  | 
|
684  | 
  "nnfis f = { y. \<exists>u x. mono_convergent u f (space M) \<and>
 | 
|
685  | 
(\<forall>n. x n \<in> psfis (u n)) \<and> x ----> y }"  | 
|
686  | 
||
| 35582 | 687  | 
lemma psfis_nnfis:  | 
688  | 
"a \<in> psfis f \<Longrightarrow> a \<in> nnfis f"  | 
|
689  | 
unfolding nnfis_def mono_convergent_def incseq_def  | 
|
690  | 
by (auto intro!: exI[of _ "\<lambda>n. f"] exI[of _ "\<lambda>n. a"] LIMSEQ_const)  | 
|
691  | 
||
| 35748 | 692  | 
lemma nnfis_0: "0 \<in> nnfis (\<lambda> x. 0)"  | 
693  | 
by (rule psfis_nnfis[OF psfis_0])  | 
|
694  | 
||
| 35582 | 695  | 
lemma nnfis_times:  | 
696  | 
assumes "a \<in> nnfis f" and "0 \<le> z"  | 
|
697  | 
shows "z * a \<in> nnfis (\<lambda>t. z * f t)"  | 
|
698  | 
proof -  | 
|
699  | 
obtain u x where "mono_convergent u f (space M)" and  | 
|
700  | 
"\<And>n. x n \<in> psfis (u n)" "x ----> a"  | 
|
701  | 
using `a \<in> nnfis f` unfolding nnfis_def by auto  | 
|
702  | 
with `0 \<le> z`show ?thesis unfolding nnfis_def mono_convergent_def incseq_def  | 
|
703  | 
by (auto intro!: exI[of _ "\<lambda>n w. z * u n w"] exI[of _ "\<lambda>n. z * x n"]  | 
|
704  | 
LIMSEQ_mult LIMSEQ_const psfis_mult mult_mono1)  | 
|
705  | 
qed  | 
|
706  | 
||
707  | 
lemma nnfis_add:  | 
|
708  | 
assumes "a \<in> nnfis f" and "b \<in> nnfis g"  | 
|
709  | 
shows "a + b \<in> nnfis (\<lambda>t. f t + g t)"  | 
|
710  | 
proof -  | 
|
711  | 
obtain u x w y  | 
|
712  | 
where "mono_convergent u f (space M)" and  | 
|
713  | 
"\<And>n. x n \<in> psfis (u n)" "x ----> a" and  | 
|
714  | 
"mono_convergent w g (space M)" and  | 
|
715  | 
"\<And>n. y n \<in> psfis (w n)" "y ----> b"  | 
|
716  | 
using `a \<in> nnfis f` `b \<in> nnfis g` unfolding nnfis_def by auto  | 
|
717  | 
thus ?thesis unfolding nnfis_def mono_convergent_def incseq_def  | 
|
718  | 
by (auto intro!: exI[of _ "\<lambda>n a. u n a + w n a"] exI[of _ "\<lambda>n. x n + y n"]  | 
|
719  | 
LIMSEQ_add LIMSEQ_const psfis_add add_mono)  | 
|
720  | 
qed  | 
|
721  | 
||
722  | 
lemma nnfis_mono:  | 
|
723  | 
assumes nnfis: "a \<in> nnfis f" "b \<in> nnfis g"  | 
|
724  | 
and mono: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"  | 
|
725  | 
shows "a \<le> b"  | 
|
726  | 
proof -  | 
|
727  | 
obtain u x w y where  | 
|
728  | 
mc: "mono_convergent u f (space M)" "mono_convergent w g (space M)" and  | 
|
729  | 
psfis: "\<And>n. x n \<in> psfis (u n)" "\<And>n. y n \<in> psfis (w n)" and  | 
|
730  | 
limseq: "x ----> a" "y ----> b" using nnfis unfolding nnfis_def by auto  | 
|
731  | 
show ?thesis  | 
|
732  | 
proof (rule LIMSEQ_le_const2[OF limseq(1)], rule exI[of _ 0], safe)  | 
|
733  | 
fix n  | 
|
734  | 
show "x n \<le> b"  | 
|
735  | 
proof (rule psfis_mono_conv_mono[OF mc(2) psfis(2) limseq(2) psfis(1)])  | 
|
736  | 
fix t assume "t \<in> space M"  | 
|
737  | 
from mono_convergent_le[OF mc(1) this] mono[OF this]  | 
|
738  | 
show "u n t \<le> g t" by (rule order_trans)  | 
|
739  | 
qed  | 
|
740  | 
qed  | 
|
741  | 
qed  | 
|
742  | 
||
743  | 
lemma nnfis_unique:  | 
|
744  | 
assumes a: "a \<in> nnfis f" and b: "b \<in> nnfis f"  | 
|
745  | 
shows "a = b"  | 
|
746  | 
using nnfis_mono[OF a b] nnfis_mono[OF b a]  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
747  | 
by (auto intro!: order_antisym[of a b])  | 
| 35582 | 748  | 
|
749  | 
lemma psfis_equiv:  | 
|
750  | 
assumes "a \<in> psfis f" and "nonneg g"  | 
|
751  | 
and "\<And>t. t \<in> space M \<Longrightarrow> f t = g t"  | 
|
752  | 
shows "a \<in> psfis g"  | 
|
753  | 
using assms unfolding psfis_def pos_simple_def by auto  | 
|
754  | 
||
755  | 
lemma psfis_mon_upclose:  | 
|
756  | 
assumes "\<And>m. a m \<in> psfis (u m)"  | 
|
757  | 
shows "\<exists>c. c \<in> psfis (mon_upclose n u)"  | 
|
758  | 
proof (induct n)  | 
|
759  | 
case 0 thus ?case unfolding mon_upclose.simps using assms ..  | 
|
760  | 
next  | 
|
761  | 
case (Suc n)  | 
|
762  | 
then obtain sn an xn where ps: "(sn, an, xn) \<in> pos_simple (mon_upclose n u)"  | 
|
763  | 
unfolding psfis_def by auto  | 
|
764  | 
obtain ss as xs where ps': "(ss, as, xs) \<in> pos_simple (u (Suc n))"  | 
|
765  | 
using assms[of "Suc n"] unfolding psfis_def by auto  | 
|
766  | 
from pos_simple_common_partition[OF ps ps'] guess x x' a s .  | 
|
767  | 
hence "(s, a, upclose x x') \<in> pos_simple (mon_upclose (Suc n) u)"  | 
|
768  | 
by (simp add: upclose_def pos_simple_def nonneg_def max_def)  | 
|
769  | 
thus ?case unfolding psfis_def by auto  | 
|
770  | 
qed  | 
|
771  | 
||
772  | 
text {* Beppo-Levi monotone convergence theorem *}
 | 
|
773  | 
lemma nnfis_mon_conv:  | 
|
774  | 
assumes mc: "mono_convergent f h (space M)"  | 
|
775  | 
and nnfis: "\<And>n. x n \<in> nnfis (f n)"  | 
|
776  | 
and "x ----> z"  | 
|
777  | 
shows "z \<in> nnfis h"  | 
|
778  | 
proof -  | 
|
779  | 
have "\<forall>n. \<exists>u y. mono_convergent u (f n) (space M) \<and> (\<forall>m. y m \<in> psfis (u m)) \<and>  | 
|
780  | 
y ----> x n"  | 
|
781  | 
using nnfis unfolding nnfis_def by auto  | 
|
782  | 
from choice[OF this] guess u ..  | 
|
783  | 
from choice[OF this] guess y ..  | 
|
784  | 
hence mc_u: "\<And>n. mono_convergent (u n) (f n) (space M)"  | 
|
785  | 
and psfis: "\<And>n m. y n m \<in> psfis (u n m)" and "\<And>n. y n ----> x n"  | 
|
786  | 
by auto  | 
|
787  | 
||
788  | 
let "?upclose n" = "mon_upclose n (\<lambda>m. u m n)"  | 
|
789  | 
||
790  | 
have "\<exists>c. \<forall>n. c n \<in> psfis (?upclose n)"  | 
|
791  | 
by (safe intro!: choice psfis_mon_upclose) (rule psfis)  | 
|
792  | 
then guess c .. note c = this[rule_format]  | 
|
793  | 
||
794  | 
show ?thesis unfolding nnfis_def  | 
|
795  | 
proof (safe intro!: exI)  | 
|
796  | 
show mc_upclose: "mono_convergent ?upclose h (space M)"  | 
|
797  | 
by (rule mon_upclose_mono_convergent[OF mc_u mc])  | 
|
798  | 
show psfis_upclose: "\<And>n. c n \<in> psfis (?upclose n)"  | 
|
799  | 
using c .  | 
|
800  | 
||
801  | 
    { fix n m :: nat assume "n \<le> m"
 | 
|
802  | 
hence "c n \<le> c m"  | 
|
803  | 
using psfis_mono[OF c c]  | 
|
804  | 
using mono_convergentD(1)[OF mc_upclose, unfolded incseq_def]  | 
|
805  | 
by auto }  | 
|
806  | 
hence "incseq c" unfolding incseq_def by auto  | 
|
807  | 
||
808  | 
    { fix n
 | 
|
809  | 
have c_nnfis: "c n \<in> nnfis (?upclose n)" using c by (rule psfis_nnfis)  | 
|
810  | 
from nnfis_mono[OF c_nnfis nnfis]  | 
|
811  | 
mon_upclose_le_mono_convergent[OF mc_u]  | 
|
812  | 
mono_convergentD(1)[OF mc]  | 
|
813  | 
have "c n \<le> x n" by fast }  | 
|
814  | 
note c_less_x = this  | 
|
815  | 
||
816  | 
    { fix n
 | 
|
817  | 
note c_less_x[of n]  | 
|
818  | 
also have "x n \<le> z"  | 
|
819  | 
proof (rule incseq_le)  | 
|
820  | 
show "x ----> z" by fact  | 
|
821  | 
from mono_convergentD(1)[OF mc]  | 
|
822  | 
show "incseq x" unfolding incseq_def  | 
|
823  | 
by (auto intro!: nnfis_mono[OF nnfis nnfis])  | 
|
824  | 
qed  | 
|
825  | 
finally have "c n \<le> z" . }  | 
|
826  | 
note c_less_z = this  | 
|
827  | 
||
828  | 
have "convergent c"  | 
|
829  | 
proof (rule Bseq_mono_convergent[unfolded incseq_def[symmetric]])  | 
|
830  | 
show "Bseq c"  | 
|
831  | 
using pos_psfis[OF c] c_less_z  | 
|
832  | 
by (auto intro!: BseqI'[of _ z])  | 
|
833  | 
show "incseq c" by fact  | 
|
834  | 
qed  | 
|
835  | 
then obtain l where l: "c ----> l" unfolding convergent_def by auto  | 
|
836  | 
||
837  | 
have "l \<le> z" using c_less_x l  | 
|
838  | 
by (auto intro!: LIMSEQ_le[OF _ `x ----> z`])  | 
|
839  | 
moreover have "z \<le> l"  | 
|
840  | 
proof (rule LIMSEQ_le_const2[OF `x ----> z`], safe intro!: exI[of _ 0])  | 
|
841  | 
fix n  | 
|
842  | 
have "l \<in> nnfis h"  | 
|
843  | 
unfolding nnfis_def using l mc_upclose psfis_upclose by auto  | 
|
844  | 
from nnfis this mono_convergent_le[OF mc]  | 
|
845  | 
show "x n \<le> l" by (rule nnfis_mono)  | 
|
846  | 
qed  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
847  | 
ultimately have "l = z" by (rule order_antisym)  | 
| 35582 | 848  | 
thus "c ----> z" using `c ----> l` by simp  | 
849  | 
qed  | 
|
850  | 
qed  | 
|
851  | 
||
852  | 
lemma nnfis_pos_on_mspace:  | 
|
853  | 
assumes "a \<in> nnfis f" and "x \<in>space M"  | 
|
854  | 
shows "0 \<le> f x"  | 
|
855  | 
using assms  | 
|
856  | 
proof -  | 
|
857  | 
from assms[unfolded nnfis_def] guess u y by auto note uy = this  | 
|
| 35748 | 858  | 
hence "\<And> n. 0 \<le> u n x"  | 
| 35582 | 859  | 
unfolding nnfis_def psfis_def pos_simple_def nonneg_def mono_convergent_def  | 
860  | 
by auto  | 
|
861  | 
thus "0 \<le> f x" using uy[rule_format]  | 
|
862  | 
unfolding nnfis_def psfis_def pos_simple_def nonneg_def mono_convergent_def  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
863  | 
using incseq_le[of "\<lambda> n. u n x" "f x"] order_trans  | 
| 35582 | 864  | 
by fast  | 
865  | 
qed  | 
|
866  | 
||
867  | 
lemma nnfis_borel_measurable:  | 
|
868  | 
assumes "a \<in> nnfis f"  | 
|
869  | 
shows "f \<in> borel_measurable M"  | 
|
870  | 
using assms unfolding nnfis_def  | 
|
871  | 
apply auto  | 
|
872  | 
apply (rule mono_convergent_borel_measurable)  | 
|
873  | 
using psfis_borel_measurable  | 
|
874  | 
by auto  | 
|
875  | 
||
876  | 
lemma borel_measurable_mon_conv_psfis:  | 
|
877  | 
assumes f_borel: "f \<in> borel_measurable M"  | 
|
878  | 
and nonneg: "\<And>t. t \<in> space M \<Longrightarrow> 0 \<le> f t"  | 
|
879  | 
shows"\<exists>u x. mono_convergent u f (space M) \<and> (\<forall>n. x n \<in> psfis (u n))"  | 
|
880  | 
proof (safe intro!: exI)  | 
|
881  | 
  let "?I n" = "{0<..<n * 2^n}"
 | 
|
882  | 
  let "?A n i" = "{w \<in> space M. real (i :: nat) / 2^(n::nat) \<le> f w \<and> f w < real (i + 1) / 2^n}"
 | 
|
883  | 
let "?u n t" = "\<Sum>i\<in>?I n. real i / 2^n * indicator_fn (?A n i) t"  | 
|
884  | 
let "?x n" = "\<Sum>i\<in>?I n. real i / 2^n * measure M (?A n i)"  | 
|
885  | 
||
886  | 
let "?w n t" = "if f t < real n then real (natfloor (f t * 2^n)) / 2^n else 0"  | 
|
887  | 
||
888  | 
  { fix t n assume t: "t \<in> space M"
 | 
|
889  | 
have "?u n t = ?w n t" (is "_ = (if _ then real ?i / _ else _)")  | 
|
890  | 
proof (cases "f t < real n")  | 
|
891  | 
case True  | 
|
892  | 
with nonneg t  | 
|
893  | 
have i: "?i < n * 2^n"  | 
|
894  | 
by (auto simp: real_of_nat_power[symmetric]  | 
|
895  | 
intro!: less_natfloor mult_nonneg_nonneg)  | 
|
896  | 
||
897  | 
hence t_in_A: "t \<in> ?A n ?i"  | 
|
898  | 
unfolding divide_le_eq less_divide_eq  | 
|
899  | 
using nonneg t True  | 
|
900  | 
by (auto intro!: real_natfloor_le  | 
|
901  | 
real_natfloor_gt_diff_one[unfolded diff_less_eq]  | 
|
902  | 
simp: real_of_nat_Suc zero_le_mult_iff)  | 
|
903  | 
||
904  | 
hence *: "real ?i / 2^n \<le> f t"  | 
|
905  | 
"f t < real (?i + 1) / 2^n" by auto  | 
|
906  | 
      { fix j assume "t \<in> ?A n j"
 | 
|
907  | 
hence "real j / 2^n \<le> f t"  | 
|
908  | 
and "f t < real (j + 1) / 2^n" by auto  | 
|
909  | 
        with * have "j \<in> {?i}" unfolding divide_le_eq less_divide_eq
 | 
|
910  | 
by auto }  | 
|
911  | 
      hence *: "\<And>j. t \<in> ?A n j \<longleftrightarrow> j \<in> {?i}" using t_in_A by auto
 | 
|
912  | 
||
913  | 
have "?u n t = real ?i / 2^n"  | 
|
914  | 
unfolding indicator_fn_def if_distrib *  | 
|
915  | 
setsum_cases[OF finite_greaterThanLessThan]  | 
|
916  | 
using i by (cases "?i = 0") simp_all  | 
|
917  | 
thus ?thesis using True by auto  | 
|
918  | 
next  | 
|
919  | 
case False  | 
|
920  | 
      have "?u n t = (\<Sum>i \<in> {0 <..< n*2^n}. 0)"
 | 
|
921  | 
proof (rule setsum_cong)  | 
|
922  | 
        fix i assume "i \<in> {0 <..< n*2^n}"
 | 
|
923  | 
hence "i + 1 \<le> n * 2^n" by simp  | 
|
924  | 
hence "real (i + 1) \<le> real n * 2^n"  | 
|
925  | 
unfolding real_of_nat_le_iff[symmetric]  | 
|
926  | 
by (auto simp: real_of_nat_power[symmetric])  | 
|
927  | 
also have "... \<le> f t * 2^n"  | 
|
928  | 
using False by (auto intro!: mult_nonneg_nonneg)  | 
|
929  | 
finally have "t \<notin> ?A n i"  | 
|
930  | 
by (auto simp: divide_le_eq less_divide_eq)  | 
|
931  | 
thus "real i / 2^n * indicator_fn (?A n i) t = 0"  | 
|
932  | 
unfolding indicator_fn_def by auto  | 
|
933  | 
qed simp  | 
|
934  | 
thus ?thesis using False by auto  | 
|
935  | 
qed }  | 
|
936  | 
note u_at_t = this  | 
|
937  | 
||
938  | 
show "mono_convergent ?u f (space M)" unfolding mono_convergent_def  | 
|
939  | 
proof safe  | 
|
940  | 
fix t assume t: "t \<in> space M"  | 
|
941  | 
    { fix m n :: nat assume "m \<le> n"
 | 
|
| 36844 | 942  | 
hence *: "(2::real)^n = 2^m * 2^(n - m)" unfolding power_add[symmetric] by auto  | 
| 35582 | 943  | 
have "real (natfloor (f t * 2^m) * natfloor (2^(n-m))) \<le> real (natfloor (f t * 2 ^ n))"  | 
944  | 
apply (subst *)  | 
|
| 36844 | 945  | 
apply (subst mult_assoc[symmetric])  | 
| 35582 | 946  | 
apply (subst real_of_nat_le_iff)  | 
947  | 
apply (rule le_mult_natfloor)  | 
|
948  | 
using nonneg[OF t] by (auto intro!: mult_nonneg_nonneg)  | 
|
949  | 
hence "real (natfloor (f t * 2^m)) * 2^n \<le> real (natfloor (f t * 2^n)) * 2^m"  | 
|
950  | 
apply (subst *)  | 
|
| 36844 | 951  | 
apply (subst (3) mult_commute)  | 
952  | 
apply (subst mult_assoc[symmetric])  | 
|
| 35582 | 953  | 
by (auto intro: mult_right_mono simp: natfloor_power real_of_nat_power[symmetric]) }  | 
954  | 
thus "incseq (\<lambda>n. ?u n t)" unfolding u_at_t[OF t] unfolding incseq_def  | 
|
955  | 
by (auto simp add: le_divide_eq divide_le_eq less_divide_eq)  | 
|
956  | 
||
957  | 
show "(\<lambda>i. ?u i t) ----> f t" unfolding u_at_t[OF t]  | 
|
958  | 
proof (rule LIMSEQ_I, safe intro!: exI)  | 
|
959  | 
fix r :: real and n :: nat  | 
|
960  | 
let ?N = "natfloor (1/r) + 1"  | 
|
961  | 
assume "0 < r" and N: "max ?N (natfloor (f t) + 1) \<le> n"  | 
|
962  | 
hence "?N \<le> n" by auto  | 
|
963  | 
||
964  | 
have "1 / r < real (natfloor (1/r) + 1)" using real_natfloor_add_one_gt  | 
|
965  | 
by (simp add: real_of_nat_Suc)  | 
|
966  | 
also have "... < 2^?N" by (rule two_realpow_gt)  | 
|
967  | 
finally have less_r: "1 / 2^?N < r" using `0 < r`  | 
|
968  | 
by (auto simp: less_divide_eq divide_less_eq algebra_simps)  | 
|
969  | 
||
970  | 
have "f t < real (natfloor (f t) + 1)" using real_natfloor_add_one_gt[of "f t"] by auto  | 
|
971  | 
also have "... \<le> real n" unfolding real_of_nat_le_iff using N by auto  | 
|
972  | 
finally have "f t < real n" .  | 
|
973  | 
||
974  | 
have "real (natfloor (f t * 2^n)) \<le> f t * 2^n"  | 
|
975  | 
using nonneg[OF t] by (auto intro!: real_natfloor_le mult_nonneg_nonneg)  | 
|
976  | 
hence less: "real (natfloor (f t * 2^n)) / 2^n \<le> f t" unfolding divide_le_eq by auto  | 
|
977  | 
||
978  | 
have "f t * 2 ^ n - 1 < real (natfloor (f t * 2^n))" using real_natfloor_gt_diff_one .  | 
|
979  | 
hence "f t - real (natfloor (f t * 2^n)) / 2^n < 1 / 2^n"  | 
|
980  | 
by (auto simp: less_divide_eq divide_less_eq algebra_simps)  | 
|
981  | 
also have "... \<le> 1 / 2^?N" using `?N \<le> n`  | 
|
982  | 
by (auto intro!: divide_left_mono mult_pos_pos simp del: power_Suc)  | 
|
983  | 
also have "... < r" using less_r .  | 
|
984  | 
finally show "norm (?w n t - f t) < r" using `f t < real n` less by auto  | 
|
985  | 
qed  | 
|
986  | 
qed  | 
|
987  | 
||
988  | 
fix n  | 
|
989  | 
show "?x n \<in> psfis (?u n)"  | 
|
990  | 
proof (rule psfis_intro)  | 
|
991  | 
show "?A n ` ?I n \<subseteq> sets M"  | 
|
992  | 
proof safe  | 
|
993  | 
fix i :: nat  | 
|
994  | 
from Int[OF  | 
|
995  | 
f_borel[unfolded borel_measurable_less_iff, rule_format, of "real (i+1) / 2^n"]  | 
|
996  | 
f_borel[unfolded borel_measurable_ge_iff, rule_format, of "real i / 2^n"]]  | 
|
997  | 
show "?A n i \<in> sets M"  | 
|
998  | 
by (metis Collect_conj_eq Int_commute Int_left_absorb Int_left_commute)  | 
|
999  | 
qed  | 
|
1000  | 
show "nonneg (\<lambda>i :: nat. real i / 2 ^ n)"  | 
|
1001  | 
unfolding nonneg_def by (auto intro!: divide_nonneg_pos)  | 
|
1002  | 
qed auto  | 
|
1003  | 
qed  | 
|
1004  | 
||
1005  | 
lemma nnfis_dom_conv:  | 
|
1006  | 
assumes borel: "f \<in> borel_measurable M"  | 
|
1007  | 
and nnfis: "b \<in> nnfis g"  | 
|
1008  | 
and ord: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"  | 
|
1009  | 
and nonneg: "\<And>t. t \<in> space M \<Longrightarrow> 0 \<le> f t"  | 
|
1010  | 
shows "\<exists>a. a \<in> nnfis f \<and> a \<le> b"  | 
|
1011  | 
proof -  | 
|
1012  | 
obtain u x where mc_f: "mono_convergent u f (space M)" and  | 
|
1013  | 
psfis: "\<And>n. x n \<in> psfis (u n)"  | 
|
1014  | 
using borel_measurable_mon_conv_psfis[OF borel nonneg] by auto  | 
|
1015  | 
||
1016  | 
  { fix n
 | 
|
1017  | 
    { fix t assume t: "t \<in> space M"
 | 
|
1018  | 
note mono_convergent_le[OF mc_f this, of n]  | 
|
1019  | 
also note ord[OF t]  | 
|
1020  | 
finally have "u n t \<le> g t" . }  | 
|
1021  | 
from nnfis_mono[OF psfis_nnfis[OF psfis] nnfis this]  | 
|
1022  | 
have "x n \<le> b" by simp }  | 
|
1023  | 
note x_less_b = this  | 
|
1024  | 
||
1025  | 
have "convergent x"  | 
|
1026  | 
proof (safe intro!: Bseq_mono_convergent)  | 
|
1027  | 
from x_less_b pos_psfis[OF psfis]  | 
|
1028  | 
show "Bseq x" by (auto intro!: BseqI'[of _ b])  | 
|
1029  | 
||
1030  | 
fix n m :: nat assume *: "n \<le> m"  | 
|
1031  | 
show "x n \<le> x m"  | 
|
1032  | 
proof (rule psfis_mono[OF `x n \<in> psfis (u n)` `x m \<in> psfis (u m)`])  | 
|
1033  | 
fix t assume "t \<in> space M"  | 
|
1034  | 
from mc_f[THEN mono_convergentD(1), unfolded incseq_def, OF this]  | 
|
1035  | 
show "u n t \<le> u m t" using * by auto  | 
|
1036  | 
qed  | 
|
1037  | 
qed  | 
|
1038  | 
then obtain a where "x ----> a" unfolding convergent_def by auto  | 
|
1039  | 
with LIMSEQ_le_const2[OF `x ----> a`] x_less_b mc_f psfis  | 
|
1040  | 
show ?thesis unfolding nnfis_def by auto  | 
|
1041  | 
qed  | 
|
1042  | 
||
1043  | 
lemma the_nnfis[simp]: "a \<in> nnfis f \<Longrightarrow> (THE a. a \<in> nnfis f) = a"  | 
|
1044  | 
by (auto intro: the_equality nnfis_unique)  | 
|
1045  | 
||
1046  | 
lemma nnfis_cong:  | 
|
1047  | 
assumes cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"  | 
|
1048  | 
shows "nnfis f = nnfis g"  | 
|
1049  | 
proof -  | 
|
1050  | 
  { fix f g :: "'a \<Rightarrow> real" assume cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"
 | 
|
1051  | 
fix x assume "x \<in> nnfis f"  | 
|
1052  | 
then guess u y unfolding nnfis_def by safe note x = this  | 
|
1053  | 
hence "mono_convergent u g (space M)"  | 
|
1054  | 
unfolding mono_convergent_def using cong by auto  | 
|
1055  | 
with x(2,3) have "x \<in> nnfis g" unfolding nnfis_def by auto }  | 
|
1056  | 
from this[OF cong] this[OF cong[symmetric]]  | 
|
1057  | 
show ?thesis by safe  | 
|
1058  | 
qed  | 
|
1059  | 
||
| 35692 | 1060  | 
section "Lebesgue Integral"  | 
1061  | 
||
1062  | 
definition  | 
|
1063  | 
"integrable f \<longleftrightarrow> (\<exists>x. x \<in> nnfis (pos_part f)) \<and> (\<exists>y. y \<in> nnfis (neg_part f))"  | 
|
1064  | 
||
1065  | 
definition  | 
|
1066  | 
"integral f = (THE i :: real. i \<in> nnfis (pos_part f)) - (THE j. j \<in> nnfis (neg_part f))"  | 
|
1067  | 
||
| 35582 | 1068  | 
lemma integral_cong:  | 
1069  | 
assumes cong: "\<And>x. x \<in> space M \<Longrightarrow> f x = g x"  | 
|
1070  | 
shows "integral f = integral g"  | 
|
1071  | 
proof -  | 
|
1072  | 
have "nnfis (pos_part f) = nnfis (pos_part g)"  | 
|
1073  | 
using cong by (auto simp: pos_part_def intro!: nnfis_cong)  | 
|
1074  | 
moreover  | 
|
1075  | 
have "nnfis (neg_part f) = nnfis (neg_part g)"  | 
|
1076  | 
using cong by (auto simp: neg_part_def intro!: nnfis_cong)  | 
|
1077  | 
ultimately show ?thesis  | 
|
1078  | 
unfolding integral_def by auto  | 
|
1079  | 
qed  | 
|
1080  | 
||
1081  | 
lemma nnfis_integral:  | 
|
1082  | 
assumes "a \<in> nnfis f"  | 
|
1083  | 
shows "integrable f" and "integral f = a"  | 
|
1084  | 
proof -  | 
|
1085  | 
have a: "a \<in> nnfis (pos_part f)"  | 
|
1086  | 
using assms nnfis_pos_on_mspace[OF assms]  | 
|
1087  | 
by (auto intro!: nnfis_mon_conv[of "\<lambda>i. f" _ "\<lambda>i. a"]  | 
|
1088  | 
LIMSEQ_const simp: mono_convergent_def pos_part_def incseq_def max_def)  | 
|
1089  | 
||
1090  | 
have "\<And>t. t \<in> space M \<Longrightarrow> neg_part f t = 0"  | 
|
1091  | 
unfolding neg_part_def min_def  | 
|
1092  | 
using nnfis_pos_on_mspace[OF assms] by auto  | 
|
1093  | 
hence 0: "0 \<in> nnfis (neg_part f)"  | 
|
1094  | 
by (auto simp: nnfis_def mono_convergent_def psfis_0 incseq_def  | 
|
1095  | 
intro!: LIMSEQ_const exI[of _ "\<lambda> x n. 0"] exI[of _ "\<lambda> n. 0"])  | 
|
1096  | 
||
1097  | 
from 0 a show "integrable f" "integral f = a"  | 
|
1098  | 
unfolding integrable_def integral_def by auto  | 
|
1099  | 
qed  | 
|
1100  | 
||
1101  | 
lemma nnfis_minus_nnfis_integral:  | 
|
1102  | 
assumes "a \<in> nnfis f" and "b \<in> nnfis g"  | 
|
1103  | 
shows "integrable (\<lambda>t. f t - g t)" and "integral (\<lambda>t. f t - g t) = a - b"  | 
|
1104  | 
proof -  | 
|
1105  | 
have borel: "(\<lambda>t. f t - g t) \<in> borel_measurable M" using assms  | 
|
1106  | 
by (blast intro:  | 
|
1107  | 
borel_measurable_diff_borel_measurable nnfis_borel_measurable)  | 
|
1108  | 
||
1109  | 
have "\<exists>x. x \<in> nnfis (pos_part (\<lambda>t. f t - g t)) \<and> x \<le> a + b"  | 
|
1110  | 
(is "\<exists>x. x \<in> nnfis ?pp \<and> _")  | 
|
1111  | 
proof (rule nnfis_dom_conv)  | 
|
1112  | 
show "?pp \<in> borel_measurable M"  | 
|
| 35692 | 1113  | 
using borel by (rule pos_part_borel_measurable neg_part_borel_measurable)  | 
| 35582 | 1114  | 
show "a + b \<in> nnfis (\<lambda>t. f t + g t)" using assms by (rule nnfis_add)  | 
1115  | 
fix t assume "t \<in> space M"  | 
|
1116  | 
with assms nnfis_add assms[THEN nnfis_pos_on_mspace[OF _ this]]  | 
|
1117  | 
show "?pp t \<le> f t + g t" unfolding pos_part_def by auto  | 
|
1118  | 
show "0 \<le> ?pp t" using nonneg_pos_part[of "\<lambda>t. f t - g t"]  | 
|
1119  | 
unfolding nonneg_def by auto  | 
|
1120  | 
qed  | 
|
1121  | 
then obtain x where x: "x \<in> nnfis ?pp" by auto  | 
|
1122  | 
moreover  | 
|
1123  | 
have "\<exists>x. x \<in> nnfis (neg_part (\<lambda>t. f t - g t)) \<and> x \<le> a + b"  | 
|
1124  | 
(is "\<exists>x. x \<in> nnfis ?np \<and> _")  | 
|
1125  | 
proof (rule nnfis_dom_conv)  | 
|
1126  | 
show "?np \<in> borel_measurable M"  | 
|
| 35692 | 1127  | 
using borel by (rule pos_part_borel_measurable neg_part_borel_measurable)  | 
| 35582 | 1128  | 
show "a + b \<in> nnfis (\<lambda>t. f t + g t)" using assms by (rule nnfis_add)  | 
1129  | 
fix t assume "t \<in> space M"  | 
|
1130  | 
with assms nnfis_add assms[THEN nnfis_pos_on_mspace[OF _ this]]  | 
|
1131  | 
show "?np t \<le> f t + g t" unfolding neg_part_def by auto  | 
|
1132  | 
show "0 \<le> ?np t" using nonneg_neg_part[of "\<lambda>t. f t - g t"]  | 
|
1133  | 
unfolding nonneg_def by auto  | 
|
1134  | 
qed  | 
|
1135  | 
then obtain y where y: "y \<in> nnfis ?np" by auto  | 
|
1136  | 
ultimately show "integrable (\<lambda>t. f t - g t)"  | 
|
1137  | 
unfolding integrable_def by auto  | 
|
1138  | 
||
1139  | 
from x and y  | 
|
1140  | 
have "a + y \<in> nnfis (\<lambda>t. f t + ?np t)"  | 
|
1141  | 
and "b + x \<in> nnfis (\<lambda>t. g t + ?pp t)" using assms by (auto intro: nnfis_add)  | 
|
1142  | 
moreover  | 
|
1143  | 
have "\<And>t. f t + ?np t = g t + ?pp t"  | 
|
1144  | 
unfolding pos_part_def neg_part_def by auto  | 
|
1145  | 
ultimately have "a - b = x - y"  | 
|
1146  | 
using nnfis_unique by (auto simp: algebra_simps)  | 
|
1147  | 
thus "integral (\<lambda>t. f t - g t) = a - b"  | 
|
1148  | 
unfolding integral_def  | 
|
1149  | 
using the_nnfis[OF x] the_nnfis[OF y] by simp  | 
|
1150  | 
qed  | 
|
1151  | 
||
1152  | 
lemma integral_borel_measurable:  | 
|
1153  | 
"integrable f \<Longrightarrow> f \<in> borel_measurable M"  | 
|
1154  | 
unfolding integrable_def  | 
|
1155  | 
by (subst pos_part_neg_part_borel_measurable_iff)  | 
|
1156  | 
(auto intro: nnfis_borel_measurable)  | 
|
1157  | 
||
1158  | 
lemma integral_indicator_fn:  | 
|
1159  | 
assumes "a \<in> sets M"  | 
|
1160  | 
shows "integral (indicator_fn a) = measure M a"  | 
|
1161  | 
and "integrable (indicator_fn a)"  | 
|
1162  | 
using psfis_indicator[OF assms, THEN psfis_nnfis]  | 
|
1163  | 
by (auto intro!: nnfis_integral)  | 
|
1164  | 
||
1165  | 
lemma integral_add:  | 
|
1166  | 
assumes "integrable f" and "integrable g"  | 
|
1167  | 
shows "integrable (\<lambda>t. f t + g t)"  | 
|
1168  | 
and "integral (\<lambda>t. f t + g t) = integral f + integral g"  | 
|
1169  | 
proof -  | 
|
1170  | 
  { fix t
 | 
|
1171  | 
have "pos_part f t + pos_part g t - (neg_part f t + neg_part g t) =  | 
|
1172  | 
f t + g t"  | 
|
1173  | 
unfolding pos_part_def neg_part_def by auto }  | 
|
1174  | 
note part_sum = this  | 
|
1175  | 
||
1176  | 
from assms obtain a b c d where  | 
|
1177  | 
a: "a \<in> nnfis (pos_part f)" and b: "b \<in> nnfis (neg_part f)" and  | 
|
1178  | 
c: "c \<in> nnfis (pos_part g)" and d: "d \<in> nnfis (neg_part g)"  | 
|
1179  | 
unfolding integrable_def by auto  | 
|
1180  | 
note sums = nnfis_add[OF a c] nnfis_add[OF b d]  | 
|
1181  | 
note int = nnfis_minus_nnfis_integral[OF sums, unfolded part_sum]  | 
|
1182  | 
||
1183  | 
show "integrable (\<lambda>t. f t + g t)" using int(1) .  | 
|
1184  | 
||
1185  | 
show "integral (\<lambda>t. f t + g t) = integral f + integral g"  | 
|
1186  | 
using int(2) sums a b c d by (simp add: the_nnfis integral_def)  | 
|
1187  | 
qed  | 
|
1188  | 
||
1189  | 
lemma integral_mono:  | 
|
1190  | 
assumes "integrable f" and "integrable g"  | 
|
1191  | 
and mono: "\<And>t. t \<in> space M \<Longrightarrow> f t \<le> g t"  | 
|
1192  | 
shows "integral f \<le> integral g"  | 
|
1193  | 
proof -  | 
|
1194  | 
from assms obtain a b c d where  | 
|
1195  | 
a: "a \<in> nnfis (pos_part f)" and b: "b \<in> nnfis (neg_part f)" and  | 
|
1196  | 
c: "c \<in> nnfis (pos_part g)" and d: "d \<in> nnfis (neg_part g)"  | 
|
1197  | 
unfolding integrable_def by auto  | 
|
1198  | 
||
1199  | 
have "a \<le> c"  | 
|
1200  | 
proof (rule nnfis_mono[OF a c])  | 
|
1201  | 
fix t assume "t \<in> space M"  | 
|
1202  | 
from mono[OF this] show "pos_part f t \<le> pos_part g t"  | 
|
1203  | 
unfolding pos_part_def by auto  | 
|
1204  | 
qed  | 
|
1205  | 
moreover have "d \<le> b"  | 
|
1206  | 
proof (rule nnfis_mono[OF d b])  | 
|
1207  | 
fix t assume "t \<in> space M"  | 
|
1208  | 
from mono[OF this] show "neg_part g t \<le> neg_part f t"  | 
|
1209  | 
unfolding neg_part_def by auto  | 
|
1210  | 
qed  | 
|
1211  | 
ultimately have "a - b \<le> c - d" by auto  | 
|
1212  | 
thus ?thesis unfolding integral_def  | 
|
1213  | 
using a b c d by (simp add: the_nnfis)  | 
|
1214  | 
qed  | 
|
1215  | 
||
1216  | 
lemma integral_uminus:  | 
|
1217  | 
assumes "integrable f"  | 
|
1218  | 
shows "integrable (\<lambda>t. - f t)"  | 
|
1219  | 
and "integral (\<lambda>t. - f t) = - integral f"  | 
|
1220  | 
proof -  | 
|
1221  | 
have "pos_part f = neg_part (\<lambda>t.-f t)" and "neg_part f = pos_part (\<lambda>t.-f t)"  | 
|
1222  | 
unfolding pos_part_def neg_part_def by (auto intro!: ext)  | 
|
1223  | 
with assms show "integrable (\<lambda>t.-f t)" and "integral (\<lambda>t.-f t) = -integral f"  | 
|
1224  | 
unfolding integrable_def integral_def by simp_all  | 
|
1225  | 
qed  | 
|
1226  | 
||
1227  | 
lemma integral_times_const:  | 
|
1228  | 
assumes "integrable f"  | 
|
1229  | 
shows "integrable (\<lambda>t. a * f t)" (is "?P a")  | 
|
1230  | 
and "integral (\<lambda>t. a * f t) = a * integral f" (is "?I a")  | 
|
1231  | 
proof -  | 
|
1232  | 
  { fix a :: real assume "0 \<le> a"
 | 
|
1233  | 
hence "pos_part (\<lambda>t. a * f t) = (\<lambda>t. a * pos_part f t)"  | 
|
1234  | 
and "neg_part (\<lambda>t. a * f t) = (\<lambda>t. a * neg_part f t)"  | 
|
1235  | 
unfolding pos_part_def neg_part_def max_def min_def  | 
|
1236  | 
by (auto intro!: ext simp: zero_le_mult_iff)  | 
|
1237  | 
moreover  | 
|
1238  | 
obtain x y where x: "x \<in> nnfis (pos_part f)" and y: "y \<in> nnfis (neg_part f)"  | 
|
1239  | 
using assms unfolding integrable_def by auto  | 
|
1240  | 
ultimately  | 
|
1241  | 
have "a * x \<in> nnfis (pos_part (\<lambda>t. a * f t))" and  | 
|
1242  | 
"a * y \<in> nnfis (neg_part (\<lambda>t. a * f t))"  | 
|
1243  | 
using nnfis_times[OF _ `0 \<le> a`] by auto  | 
|
1244  | 
with x y have "?P a \<and> ?I a"  | 
|
1245  | 
unfolding integrable_def integral_def by (auto simp: algebra_simps) }  | 
|
1246  | 
note int = this  | 
|
1247  | 
||
1248  | 
have "?P a \<and> ?I a"  | 
|
1249  | 
proof (cases "0 \<le> a")  | 
|
1250  | 
case True from int[OF this] show ?thesis .  | 
|
1251  | 
next  | 
|
1252  | 
case False with int[of "- a"] integral_uminus[of "\<lambda>t. - a * f t"]  | 
|
1253  | 
show ?thesis by auto  | 
|
1254  | 
qed  | 
|
1255  | 
thus "integrable (\<lambda>t. a * f t)"  | 
|
1256  | 
and "integral (\<lambda>t. a * f t) = a * integral f" by simp_all  | 
|
1257  | 
qed  | 
|
1258  | 
||
1259  | 
lemma integral_cmul_indicator:  | 
|
1260  | 
assumes "s \<in> sets M"  | 
|
1261  | 
shows "integral (\<lambda>x. c * indicator_fn s x) = c * (measure M s)"  | 
|
1262  | 
and "integrable (\<lambda>x. c * indicator_fn s x)"  | 
|
1263  | 
using assms integral_times_const integral_indicator_fn by auto  | 
|
1264  | 
||
1265  | 
lemma integral_zero:  | 
|
1266  | 
shows "integral (\<lambda>x. 0) = 0"  | 
|
1267  | 
and "integrable (\<lambda>x. 0)"  | 
|
1268  | 
using integral_cmul_indicator[OF empty_sets, of 0]  | 
|
1269  | 
unfolding indicator_fn_def by auto  | 
|
1270  | 
||
1271  | 
lemma integral_setsum:  | 
|
1272  | 
assumes "finite S"  | 
|
1273  | 
assumes "\<And>n. n \<in> S \<Longrightarrow> integrable (f n)"  | 
|
1274  | 
shows "integral (\<lambda>x. \<Sum> i \<in> S. f i x) = (\<Sum> i \<in> S. integral (f i))" (is "?int S")  | 
|
1275  | 
and "integrable (\<lambda>x. \<Sum> i \<in> S. f i x)" (is "?I s")  | 
|
1276  | 
proof -  | 
|
1277  | 
from assms have "?int S \<and> ?I S"  | 
|
1278  | 
proof (induct S)  | 
|
1279  | 
case empty thus ?case by (simp add: integral_zero)  | 
|
1280  | 
next  | 
|
1281  | 
case (insert i S)  | 
|
1282  | 
thus ?case  | 
|
1283  | 
apply simp  | 
|
1284  | 
apply (subst integral_add)  | 
|
1285  | 
using assms apply auto  | 
|
1286  | 
apply (subst integral_add)  | 
|
1287  | 
using assms by auto  | 
|
1288  | 
qed  | 
|
1289  | 
thus "?int S" and "?I S" by auto  | 
|
1290  | 
qed  | 
|
1291  | 
||
| 36624 | 1292  | 
lemma (in measure_space) integrable_abs:  | 
1293  | 
assumes "integrable f"  | 
|
1294  | 
shows "integrable (\<lambda> x. \<bar>f x\<bar>)"  | 
|
1295  | 
using assms  | 
|
1296  | 
proof -  | 
|
1297  | 
from assms obtain p q where pq: "p \<in> nnfis (pos_part f)" "q \<in> nnfis (neg_part f)"  | 
|
1298  | 
unfolding integrable_def by auto  | 
|
1299  | 
hence "p + q \<in> nnfis (\<lambda> x. pos_part f x + neg_part f x)"  | 
|
1300  | 
using nnfis_add by auto  | 
|
1301  | 
hence "p + q \<in> nnfis (\<lambda> x. \<bar>f x\<bar>)" using pos_neg_part_abs[of f] by simp  | 
|
1302  | 
thus ?thesis unfolding integrable_def  | 
|
1303  | 
using ext[OF pos_part_abs[of f], of "\<lambda> y. y"]  | 
|
1304  | 
ext[OF neg_part_abs[of f], of "\<lambda> y. y"]  | 
|
1305  | 
using nnfis_0 by auto  | 
|
1306  | 
qed  | 
|
1307  | 
||
| 35582 | 1308  | 
lemma markov_ineq:  | 
1309  | 
assumes "integrable f" "0 < a" "integrable (\<lambda>x. \<bar>f x\<bar>^n)"  | 
|
1310  | 
  shows "measure M (f -` {a ..} \<inter> space M) \<le> integral (\<lambda>x. \<bar>f x\<bar>^n) / a^n"
 | 
|
1311  | 
using assms  | 
|
1312  | 
proof -  | 
|
1313  | 
from assms have "0 < a ^ n" using real_root_pow_pos by auto  | 
|
1314  | 
from assms have "f \<in> borel_measurable M"  | 
|
1315  | 
using integral_borel_measurable[OF `integrable f`] by auto  | 
|
1316  | 
  hence w: "{w . w \<in> space M \<and> a \<le> f w} \<in> sets M"
 | 
|
1317  | 
using borel_measurable_ge_iff by auto  | 
|
1318  | 
  have i: "integrable (indicator_fn {w . w \<in> space M \<and> a \<le> f w})"
 | 
|
1319  | 
using integral_indicator_fn[OF w] by simp  | 
|
1320  | 
  have v1: "\<And> t. a ^ n * (indicator_fn {w . w \<in> space M \<and> a \<le> f w}) t 
 | 
|
1321  | 
            \<le> (f t) ^ n * (indicator_fn {w . w \<in> space M \<and> a \<le> f w}) t"
 | 
|
1322  | 
unfolding indicator_fn_def  | 
|
1323  | 
using `0 < a` power_mono[of a] assms by auto  | 
|
1324  | 
  have v2: "\<And> t. (f t) ^ n * (indicator_fn {w . w \<in> space M \<and> a \<le> f w}) t \<le> \<bar>f t\<bar> ^ n"
 | 
|
1325  | 
unfolding indicator_fn_def  | 
|
1326  | 
using power_mono[of a _ n] abs_ge_self `a > 0`  | 
|
1327  | 
by auto  | 
|
1328  | 
  have "{w \<in> space M. a \<le> f w} \<inter> space M = {w . a \<le> f w} \<inter> space M"
 | 
|
1329  | 
using Collect_eq by auto  | 
|
1330  | 
from Int_absorb2[OF sets_into_space[OF w]] `0 < a ^ n` sets_into_space[OF w] w this  | 
|
1331  | 
  have "(a ^ n) * (measure M ((f -` {y . a \<le> y}) \<inter> space M)) =
 | 
|
1332  | 
        (a ^ n) * measure M {w . w \<in> space M \<and> a \<le> f w}"
 | 
|
1333  | 
unfolding vimage_Collect_eq[of f] by simp  | 
|
1334  | 
  also have "\<dots> = integral (\<lambda> t. a ^ n * (indicator_fn {w . w \<in> space M \<and> a \<le> f w}) t)"
 | 
|
1335  | 
using integral_cmul_indicator[OF w] i by auto  | 
|
1336  | 
also have "\<dots> \<le> integral (\<lambda> t. \<bar> f t \<bar> ^ n)"  | 
|
1337  | 
apply (rule integral_mono)  | 
|
1338  | 
using integral_cmul_indicator[OF w]  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
1339  | 
`integrable (\<lambda> x. \<bar>f x\<bar> ^ n)` order_trans[OF v1 v2] by auto  | 
| 35582 | 1340  | 
  finally show "measure M (f -` {a ..} \<inter> space M) \<le> integral (\<lambda>x. \<bar>f x\<bar>^n) / a^n"
 | 
1341  | 
unfolding atLeast_def  | 
|
| 
36778
 
739a9379e29b
avoid using real-specific versions of generic lemmas
 
huffman 
parents: 
36725 
diff
changeset
 | 
1342  | 
by (auto intro!: mult_imp_le_div_pos[OF `0 < a ^ n`], simp add: mult_commute)  | 
| 35582 | 1343  | 
qed  | 
1344  | 
||
| 36624 | 1345  | 
lemma (in measure_space) integral_0:  | 
1346  | 
fixes f :: "'a \<Rightarrow> real"  | 
|
1347  | 
assumes "integrable f" "integral f = 0" "nonneg f" and borel: "f \<in> borel_measurable M"  | 
|
1348  | 
  shows "measure M ({x. f x \<noteq> 0} \<inter> space M) = 0"
 | 
|
1349  | 
proof -  | 
|
1350  | 
  have "{x. f x \<noteq> 0} = {x. \<bar>f x\<bar> > 0}" by auto
 | 
|
1351  | 
moreover  | 
|
1352  | 
  { fix y assume "y \<in> {x. \<bar> f x \<bar> > 0}"
 | 
|
1353  | 
hence "\<bar> f y \<bar> > 0" by auto  | 
|
1354  | 
hence "\<exists> n. \<bar>f y\<bar> \<ge> inverse (real (Suc n))"  | 
|
1355  | 
using ex_inverse_of_nat_Suc_less[of "\<bar>f y\<bar>"] less_imp_le unfolding real_of_nat_def by auto  | 
|
1356  | 
    hence "y \<in> (\<Union> n. {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))})"
 | 
|
1357  | 
by auto }  | 
|
1358  | 
moreover  | 
|
1359  | 
  { fix y assume "y \<in> (\<Union> n. {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))})"
 | 
|
1360  | 
    then obtain n where n: "y \<in> {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))}" by auto
 | 
|
1361  | 
hence "\<bar>f y\<bar> \<ge> inverse (real (Suc n))" by auto  | 
|
1362  | 
hence "\<bar>f y\<bar> > 0"  | 
|
1363  | 
using real_of_nat_Suc_gt_zero  | 
|
1364  | 
positive_imp_inverse_positive[of "real_of_nat (Suc n)"] by fastsimp  | 
|
1365  | 
    hence "y \<in> {x. \<bar>f x\<bar> > 0}" by auto }
 | 
|
1366  | 
  ultimately have fneq0_UN: "{x. f x \<noteq> 0} = (\<Union> n. {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))})"
 | 
|
1367  | 
by blast  | 
|
1368  | 
  { fix n
 | 
|
1369  | 
have int_one: "integrable (\<lambda> x. \<bar>f x\<bar> ^ 1)" using integrable_abs assms by auto  | 
|
1370  | 
    have "measure M (f -` {inverse (real (Suc n))..} \<inter> space M)
 | 
|
1371  | 
\<le> integral (\<lambda> x. \<bar>f x\<bar> ^ 1) / (inverse (real (Suc n)) ^ 1)"  | 
|
1372  | 
using markov_ineq[OF `integrable f` _ int_one] real_of_nat_Suc_gt_zero by auto  | 
|
1373  | 
    hence le0: "measure M (f -` {inverse (real (Suc n))..} \<inter> space M) \<le> 0"
 | 
|
1374  | 
using assms unfolding nonneg_def by auto  | 
|
1375  | 
    have "{x. f x \<ge> inverse (real (Suc n))} \<inter> space M \<in> sets M"
 | 
|
1376  | 
apply (subst Int_commute) unfolding Int_def  | 
|
1377  | 
using borel[unfolded borel_measurable_ge_iff] by simp  | 
|
1378  | 
    hence m0: "measure M ({x. f x \<ge> inverse (real (Suc n))} \<inter> space M) = 0 \<and>
 | 
|
1379  | 
      {x. f x \<ge> inverse (real (Suc n))} \<inter> space M \<in> sets M"
 | 
|
1380  | 
using positive le0 unfolding atLeast_def by fastsimp }  | 
|
1381  | 
  moreover hence "range (\<lambda> n. {x. f x \<ge> inverse (real (Suc n))} \<inter> space M) \<subseteq> sets M"
 | 
|
1382  | 
by auto  | 
|
1383  | 
moreover  | 
|
1384  | 
  { fix n
 | 
|
1385  | 
have "inverse (real (Suc n)) \<ge> inverse (real (Suc (Suc n)))"  | 
|
1386  | 
using less_imp_inverse_less real_of_nat_Suc_gt_zero[of n] by fastsimp  | 
|
1387  | 
hence "\<And> x. f x \<ge> inverse (real (Suc n)) \<Longrightarrow> f x \<ge> inverse (real (Suc (Suc n)))" by (rule order_trans)  | 
|
1388  | 
    hence "{x. f x \<ge> inverse (real (Suc n))} \<inter> space M
 | 
|
1389  | 
         \<subseteq> {x. f x \<ge> inverse (real (Suc (Suc n)))} \<inter> space M" by auto }
 | 
|
1390  | 
  ultimately have "(\<lambda> x. 0) ----> measure M (\<Union> n. {x. f x \<ge> inverse (real (Suc n))} \<inter> space M)"
 | 
|
1391  | 
    using monotone_convergence[of "\<lambda> n. {x. f x \<ge> inverse (real (Suc n))} \<inter> space M"]
 | 
|
1392  | 
unfolding o_def by (simp del: of_nat_Suc)  | 
|
1393  | 
  hence "measure M (\<Union> n. {x. f x \<ge> inverse (real (Suc n))} \<inter> space M) = 0"
 | 
|
1394  | 
using LIMSEQ_const[of 0] LIMSEQ_unique by simp  | 
|
1395  | 
  hence "measure M ((\<Union> n. {x. \<bar>f x\<bar> \<ge> inverse (real (Suc n))}) \<inter> space M) = 0"
 | 
|
1396  | 
using assms unfolding nonneg_def by auto  | 
|
1397  | 
  thus "measure M ({x. f x \<noteq> 0} \<inter> space M) = 0" using fneq0_UN by simp
 | 
|
1398  | 
qed  | 
|
1399  | 
||
| 35748 | 1400  | 
section "Lebesgue integration on countable spaces"  | 
1401  | 
||
1402  | 
lemma nnfis_on_countable:  | 
|
1403  | 
assumes borel: "f \<in> borel_measurable M"  | 
|
1404  | 
  and bij: "bij_betw enum S (f ` space M - {0})"
 | 
|
1405  | 
  and enum_zero: "enum ` (-S) \<subseteq> {0}"
 | 
|
1406  | 
and nn_enum: "\<And>n. 0 \<le> enum n"  | 
|
1407  | 
  and sums: "(\<lambda>r. enum r * measure M (f -` {enum r} \<inter> space M)) sums x" (is "?sum sums x")
 | 
|
1408  | 
shows "x \<in> nnfis f"  | 
|
1409  | 
proof -  | 
|
1410  | 
have inj_enum: "inj_on enum S"  | 
|
1411  | 
    and range_enum: "enum ` S = f ` space M - {0}"
 | 
|
1412  | 
using bij by (auto simp: bij_betw_def)  | 
|
1413  | 
||
1414  | 
  let "?x n z" = "\<Sum>i = 0..<n. enum i * indicator_fn (f -` {enum i} \<inter> space M) z"
 | 
|
1415  | 
||
1416  | 
show ?thesis  | 
|
1417  | 
proof (rule nnfis_mon_conv)  | 
|
1418  | 
show "(\<lambda>n. \<Sum>i = 0..<n. ?sum i) ----> x" using sums unfolding sums_def .  | 
|
1419  | 
next  | 
|
1420  | 
fix n  | 
|
1421  | 
show "(\<Sum>i = 0..<n. ?sum i) \<in> nnfis (?x n)"  | 
|
1422  | 
proof (induct n)  | 
|
1423  | 
case 0 thus ?case by (simp add: nnfis_0)  | 
|
1424  | 
next  | 
|
1425  | 
case (Suc n) thus ?case using nn_enum  | 
|
1426  | 
by (auto intro!: nnfis_add nnfis_times psfis_nnfis[OF psfis_indicator] borel_measurable_vimage[OF borel])  | 
|
1427  | 
qed  | 
|
1428  | 
next  | 
|
1429  | 
show "mono_convergent ?x f (space M)"  | 
|
1430  | 
proof (rule mono_convergentI)  | 
|
1431  | 
fix x  | 
|
1432  | 
show "incseq (\<lambda>n. ?x n x)"  | 
|
1433  | 
by (rule incseq_SucI, auto simp: indicator_fn_def nn_enum)  | 
|
1434  | 
||
1435  | 
      have fin: "\<And>n. finite (enum ` ({0..<n} \<inter> S))" by auto
 | 
|
1436  | 
||
1437  | 
assume "x \<in> space M"  | 
|
1438  | 
hence "f x \<in> enum ` S \<or> f x = 0" using range_enum by auto  | 
|
1439  | 
thus "(\<lambda>n. ?x n x) ----> f x"  | 
|
1440  | 
proof (rule disjE)  | 
|
1441  | 
assume "f x \<in> enum ` S"  | 
|
1442  | 
then obtain i where "i \<in> S" and "f x = enum i" by auto  | 
|
1443  | 
||
1444  | 
        { fix n
 | 
|
1445  | 
have sum_ranges:  | 
|
1446  | 
            "i < n \<Longrightarrow> enum`({0..<n} \<inter> S) \<inter> {z. enum i = z \<and> x\<in>space M} = {enum i}"
 | 
|
1447  | 
            "\<not> i < n \<Longrightarrow> enum`({0..<n} \<inter> S) \<inter> {z. enum i = z \<and> x\<in>space M} = {}"
 | 
|
1448  | 
using `x \<in> space M` `i \<in> S` inj_enum[THEN inj_on_iff] by auto  | 
|
1449  | 
have "?x n x =  | 
|
1450  | 
            (\<Sum>i \<in> {0..<n} \<inter> S. enum i * indicator_fn (f -` {enum i} \<inter> space M) x)"
 | 
|
1451  | 
using enum_zero by (auto intro!: setsum_mono_zero_cong_right)  | 
|
1452  | 
also have "... =  | 
|
1453  | 
            (\<Sum>z \<in> enum`({0..<n} \<inter> S). z * indicator_fn (f -` {z} \<inter> space M) x)"
 | 
|
1454  | 
using inj_enum[THEN subset_inj_on] by (auto simp: setsum_reindex)  | 
|
1455  | 
also have "... = (if i < n then f x else 0)"  | 
|
1456  | 
unfolding indicator_fn_def if_distrib[where x=1 and y=0]  | 
|
1457  | 
setsum_cases[OF fin]  | 
|
1458  | 
using sum_ranges `f x = enum i`  | 
|
1459  | 
by auto  | 
|
1460  | 
finally have "?x n x = (if i < n then f x else 0)" . }  | 
|
1461  | 
note sum_equals_if = this  | 
|
1462  | 
||
1463  | 
show ?thesis unfolding sum_equals_if  | 
|
1464  | 
by (rule LIMSEQ_offset[where k="i + 1"]) (auto intro!: LIMSEQ_const)  | 
|
1465  | 
next  | 
|
1466  | 
assume "f x = 0"  | 
|
1467  | 
        { fix n have "?x n x = 0"
 | 
|
1468  | 
unfolding indicator_fn_def if_distrib[where x=1 and y=0]  | 
|
1469  | 
setsum_cases[OF finite_atLeastLessThan]  | 
|
1470  | 
using `f x = 0` `x \<in> space M`  | 
|
1471  | 
by (auto split: split_if) }  | 
|
1472  | 
thus ?thesis using `f x = 0` by (auto intro!: LIMSEQ_const)  | 
|
1473  | 
qed  | 
|
1474  | 
qed  | 
|
1475  | 
qed  | 
|
1476  | 
qed  | 
|
1477  | 
||
1478  | 
lemma integral_on_countable:  | 
|
| 35833 | 1479  | 
fixes enum :: "nat \<Rightarrow> real"  | 
| 35748 | 1480  | 
assumes borel: "f \<in> borel_measurable M"  | 
1481  | 
and bij: "bij_betw enum S (f ` space M)"  | 
|
1482  | 
  and enum_zero: "enum ` (-S) \<subseteq> {0}"
 | 
|
1483  | 
  and abs_summable: "summable (\<lambda>r. \<bar>enum r * measure M (f -` {enum r} \<inter> space M)\<bar>)"
 | 
|
1484  | 
shows "integrable f"  | 
|
1485  | 
  and "integral f = (\<Sum>r. enum r * measure M (f -` {enum r} \<inter> space M))" (is "_ = suminf (?sum f enum)")
 | 
|
1486  | 
proof -  | 
|
1487  | 
  { fix f enum
 | 
|
1488  | 
assume borel: "f \<in> borel_measurable M"  | 
|
1489  | 
and bij: "bij_betw enum S (f ` space M)"  | 
|
1490  | 
      and enum_zero: "enum ` (-S) \<subseteq> {0}"
 | 
|
1491  | 
      and abs_summable: "summable (\<lambda>r. \<bar>enum r * measure M (f -` {enum r} \<inter> space M)\<bar>)"
 | 
|
1492  | 
have inj_enum: "inj_on enum S" and range_enum: "f ` space M = enum ` S"  | 
|
1493  | 
using bij unfolding bij_betw_def by auto  | 
|
1494  | 
||
1495  | 
    have [simp, intro]: "\<And>X. 0 \<le> measure M (f -` {X} \<inter> space M)"
 | 
|
1496  | 
by (rule positive, rule borel_measurable_vimage[OF borel])  | 
|
1497  | 
||
1498  | 
have "(\<Sum>r. ?sum (pos_part f) (pos_part enum) r) \<in> nnfis (pos_part f) \<and>  | 
|
1499  | 
summable (\<lambda>r. ?sum (pos_part f) (pos_part enum) r)"  | 
|
1500  | 
proof (rule conjI, rule nnfis_on_countable)  | 
|
1501  | 
      have pos_f_image: "pos_part f ` space M - {0} = f ` space M \<inter> {0<..}"
 | 
|
1502  | 
unfolding pos_part_def max_def by auto  | 
|
1503  | 
||
1504  | 
      show "bij_betw (pos_part enum) {x \<in> S. 0 < enum x} (pos_part f ` space M - {0})"
 | 
|
1505  | 
unfolding bij_betw_def pos_f_image  | 
|
1506  | 
unfolding pos_part_def max_def range_enum  | 
|
1507  | 
by (auto intro!: inj_onI simp: inj_enum[THEN inj_on_eq_iff])  | 
|
1508  | 
||
1509  | 
show "\<And>n. 0 \<le> pos_part enum n" unfolding pos_part_def max_def by auto  | 
|
1510  | 
||
1511  | 
show "pos_part f \<in> borel_measurable M"  | 
|
1512  | 
by (rule pos_part_borel_measurable[OF borel])  | 
|
1513  | 
||
1514  | 
      show "pos_part enum ` (- {x \<in> S. 0 < enum x}) \<subseteq> {0}"
 | 
|
1515  | 
unfolding pos_part_def max_def using enum_zero by auto  | 
|
1516  | 
||
1517  | 
show "summable (\<lambda>r. ?sum (pos_part f) (pos_part enum) r)"  | 
|
1518  | 
proof (rule summable_comparison_test[OF _ abs_summable], safe intro!: exI[of _ 0])  | 
|
1519  | 
fix n :: nat  | 
|
1520  | 
        have "pos_part enum n \<noteq> 0 \<Longrightarrow> (pos_part f -` {enum n} \<inter> space M) =
 | 
|
1521  | 
          (if 0 < enum n then (f -` {enum n} \<inter> space M) else {})"
 | 
|
1522  | 
unfolding pos_part_def max_def by (auto split: split_if_asm)  | 
|
1523  | 
thus "norm (?sum (pos_part f) (pos_part enum) n) \<le> \<bar>?sum f enum n \<bar>"  | 
|
1524  | 
by (cases "pos_part enum n = 0",  | 
|
1525  | 
auto simp: pos_part_def max_def abs_mult not_le split: split_if_asm intro!: mult_nonpos_nonneg)  | 
|
1526  | 
qed  | 
|
1527  | 
thus "(\<lambda>r. ?sum (pos_part f) (pos_part enum) r) sums (\<Sum>r. ?sum (pos_part f) (pos_part enum) r)"  | 
|
1528  | 
by (rule summable_sums)  | 
|
1529  | 
qed }  | 
|
1530  | 
note pos = this  | 
|
1531  | 
||
1532  | 
note pos_part = pos[OF assms(1-4)]  | 
|
1533  | 
moreover  | 
|
1534  | 
have neg_part_to_pos_part:  | 
|
1535  | 
"\<And>f :: _ \<Rightarrow> real. neg_part f = pos_part (uminus \<circ> f)"  | 
|
1536  | 
by (auto simp: pos_part_def neg_part_def min_def max_def expand_fun_eq)  | 
|
1537  | 
||
1538  | 
have neg_part: "(\<Sum>r. ?sum (neg_part f) (neg_part enum) r) \<in> nnfis (neg_part f) \<and>  | 
|
1539  | 
summable (\<lambda>r. ?sum (neg_part f) (neg_part enum) r)"  | 
|
1540  | 
unfolding neg_part_to_pos_part  | 
|
1541  | 
proof (rule pos)  | 
|
1542  | 
show "uminus \<circ> f \<in> borel_measurable M" unfolding comp_def  | 
|
1543  | 
by (rule borel_measurable_uminus_borel_measurable[OF borel])  | 
|
1544  | 
||
1545  | 
show "bij_betw (uminus \<circ> enum) S ((uminus \<circ> f) ` space M)"  | 
|
1546  | 
using bij unfolding bij_betw_def  | 
|
1547  | 
by (auto intro!: comp_inj_on simp: image_compose)  | 
|
1548  | 
||
1549  | 
    show "(uminus \<circ> enum) ` (- S) \<subseteq> {0}"
 | 
|
1550  | 
using enum_zero by auto  | 
|
1551  | 
||
1552  | 
    have minus_image: "\<And>r. (uminus \<circ> f) -` {(uminus \<circ> enum) r} \<inter> space M = f -` {enum r} \<inter> space M"
 | 
|
1553  | 
by auto  | 
|
1554  | 
    show "summable (\<lambda>r. \<bar>(uminus \<circ> enum) r * measure_space.measure M ((uminus \<circ> f) -` {(uminus \<circ> enum) r} \<inter> space M)\<bar>)"
 | 
|
1555  | 
unfolding minus_image using abs_summable by simp  | 
|
1556  | 
qed  | 
|
1557  | 
ultimately show "integrable f" unfolding integrable_def by auto  | 
|
1558  | 
||
1559  | 
  { fix r
 | 
|
1560  | 
have "?sum (pos_part f) (pos_part enum) r - ?sum (neg_part f) (neg_part enum) r = ?sum f enum r"  | 
|
1561  | 
proof (cases rule: linorder_cases)  | 
|
1562  | 
assume "0 < enum r"  | 
|
1563  | 
      hence "pos_part f -` {enum r} \<inter> space M = f -` {enum r} \<inter> space M"
 | 
|
1564  | 
unfolding pos_part_def max_def by (auto split: split_if_asm)  | 
|
1565  | 
with `0 < enum r` show ?thesis unfolding pos_part_def neg_part_def min_def max_def expand_fun_eq  | 
|
1566  | 
by auto  | 
|
1567  | 
next  | 
|
1568  | 
assume "enum r < 0"  | 
|
1569  | 
      hence "neg_part f -` {- enum r} \<inter> space M = f -` {enum r} \<inter> space M"
 | 
|
1570  | 
unfolding neg_part_def min_def by (auto split: split_if_asm)  | 
|
1571  | 
with `enum r < 0` show ?thesis unfolding pos_part_def neg_part_def min_def max_def expand_fun_eq  | 
|
1572  | 
by auto  | 
|
1573  | 
qed (simp add: neg_part_def pos_part_def) }  | 
|
1574  | 
note sum_diff_eq_sum = this  | 
|
1575  | 
||
1576  | 
have "(\<Sum>r. ?sum (pos_part f) (pos_part enum) r) - (\<Sum>r. ?sum (neg_part f) (neg_part enum) r)  | 
|
1577  | 
= (\<Sum>r. ?sum (pos_part f) (pos_part enum) r - ?sum (neg_part f) (neg_part enum) r)"  | 
|
1578  | 
using neg_part pos_part by (auto intro: suminf_diff)  | 
|
1579  | 
also have "... = (\<Sum>r. ?sum f enum r)" unfolding sum_diff_eq_sum ..  | 
|
1580  | 
finally show "integral f = suminf (?sum f enum)"  | 
|
1581  | 
unfolding integral_def using pos_part neg_part  | 
|
1582  | 
by (auto dest: the_nnfis)  | 
|
1583  | 
qed  | 
|
1584  | 
||
| 35692 | 1585  | 
section "Lebesgue integration on finite space"  | 
1586  | 
||
| 35582 | 1587  | 
lemma integral_finite_on_sets:  | 
1588  | 
assumes "f \<in> borel_measurable M" and "finite (space M)" and "a \<in> sets M"  | 
|
1589  | 
shows "integral (\<lambda>x. f x * indicator_fn a x) =  | 
|
1590  | 
    (\<Sum> r \<in> f`a. r * measure M (f -` {r} \<inter> a))" (is "integral ?f = _")
 | 
|
1591  | 
proof -  | 
|
1592  | 
  { fix x assume "x \<in> a"
 | 
|
1593  | 
    with assms have "f -` {f x} \<inter> space M \<in> sets M"
 | 
|
1594  | 
by (subst Int_commute)  | 
|
1595  | 
(auto simp: vimage_def Int_def  | 
|
1596  | 
intro!: borel_measurable_const  | 
|
1597  | 
borel_measurable_eq_borel_measurable)  | 
|
1598  | 
from Int[OF this assms(3)]  | 
|
1599  | 
sets_into_space[OF assms(3), THEN Int_absorb1]  | 
|
1600  | 
    have "f -` {f x} \<inter> a \<in> sets M" by (simp add: Int_assoc) }
 | 
|
1601  | 
note vimage_f = this  | 
|
1602  | 
||
1603  | 
have "finite a"  | 
|
1604  | 
using assms(2,3) sets_into_space  | 
|
1605  | 
by (auto intro: finite_subset)  | 
|
1606  | 
||
1607  | 
have "integral (\<lambda>x. f x * indicator_fn a x) =  | 
|
1608  | 
    integral (\<lambda>x. \<Sum>i\<in>f ` a. i * indicator_fn (f -` {i} \<inter> a) x)"
 | 
|
1609  | 
(is "_ = integral (\<lambda>x. setsum (?f x) _)")  | 
|
1610  | 
unfolding indicator_fn_def if_distrib  | 
|
1611  | 
using `finite a` by (auto simp: setsum_cases intro!: integral_cong)  | 
|
1612  | 
also have "\<dots> = (\<Sum>i\<in>f`a. integral (\<lambda>x. ?f x i))"  | 
|
1613  | 
proof (rule integral_setsum, safe)  | 
|
1614  | 
fix n x assume "x \<in> a"  | 
|
1615  | 
thus "integrable (\<lambda>y. ?f y (f x))"  | 
|
1616  | 
using integral_indicator_fn(2)[OF vimage_f]  | 
|
1617  | 
by (auto intro!: integral_times_const)  | 
|
1618  | 
qed (simp add: `finite a`)  | 
|
1619  | 
  also have "\<dots> = (\<Sum>i\<in>f`a. i * measure M (f -` {i} \<inter> a))"
 | 
|
1620  | 
using integral_cmul_indicator[OF vimage_f]  | 
|
1621  | 
by (auto intro!: setsum_cong)  | 
|
1622  | 
finally show ?thesis .  | 
|
1623  | 
qed  | 
|
1624  | 
||
1625  | 
lemma integral_finite:  | 
|
1626  | 
assumes "f \<in> borel_measurable M" and "finite (space M)"  | 
|
1627  | 
  shows "integral f = (\<Sum> r \<in> f ` space M. r * measure M (f -` {r} \<inter> space M))"
 | 
|
1628  | 
using integral_finite_on_sets[OF assms top]  | 
|
1629  | 
integral_cong[of "\<lambda>x. f x * indicator_fn (space M) x" f]  | 
|
1630  | 
by (auto simp add: indicator_fn_def)  | 
|
1631  | 
||
| 35692 | 1632  | 
section "Radon–Nikodym derivative"  | 
| 35582 | 1633  | 
|
| 35692 | 1634  | 
definition  | 
1635  | 
"RN_deriv v \<equiv> SOME f. measure_space (M\<lparr>measure := v\<rparr>) \<and>  | 
|
1636  | 
f \<in> borel_measurable M \<and>  | 
|
1637  | 
(\<forall>a \<in> sets M. (integral (\<lambda>x. f x * indicator_fn a x) = v a))"  | 
|
| 35582 | 1638  | 
|
| 35977 | 1639  | 
end  | 
1640  | 
||
1641  | 
lemma sigma_algebra_cong:  | 
|
1642  | 
  fixes M :: "('a, 'b) algebra_scheme" and M' :: "('a, 'c) algebra_scheme"
 | 
|
1643  | 
assumes *: "sigma_algebra M"  | 
|
1644  | 
and cong: "space M = space M'" "sets M = sets M'"  | 
|
1645  | 
shows "sigma_algebra M'"  | 
|
1646  | 
using * unfolding sigma_algebra_def algebra_def sigma_algebra_axioms_def unfolding cong .  | 
|
1647  | 
||
1648  | 
lemma finite_Pow_additivity_sufficient:  | 
|
1649  | 
assumes "finite (space M)" and "sets M = Pow (space M)"  | 
|
1650  | 
and "positive M (measure M)" and "additive M (measure M)"  | 
|
1651  | 
shows "finite_measure_space M"  | 
|
1652  | 
proof -  | 
|
1653  | 
have "sigma_algebra M"  | 
|
1654  | 
using assms by (auto intro!: sigma_algebra_cong[OF sigma_algebra_Pow])  | 
|
1655  | 
||
1656  | 
have "measure_space M"  | 
|
1657  | 
by (rule Measure.finite_additivity_sufficient) (fact+)  | 
|
1658  | 
thus ?thesis  | 
|
1659  | 
unfolding finite_measure_space_def finite_measure_space_axioms_def  | 
|
1660  | 
using assms by simp  | 
|
1661  | 
qed  | 
|
1662  | 
||
1663  | 
lemma finite_measure_spaceI:  | 
|
1664  | 
assumes "measure_space M" and "finite (space M)" and "sets M = Pow (space M)"  | 
|
1665  | 
shows "finite_measure_space M"  | 
|
1666  | 
unfolding finite_measure_space_def finite_measure_space_axioms_def  | 
|
1667  | 
using assms by simp  | 
|
1668  | 
||
1669  | 
lemma (in finite_measure_space) integral_finite_singleton:  | 
|
1670  | 
  "integral f = (\<Sum>x \<in> space M. f x * measure M {x})"
 | 
|
1671  | 
proof -  | 
|
1672  | 
have "f \<in> borel_measurable M"  | 
|
1673  | 
unfolding borel_measurable_le_iff  | 
|
1674  | 
using sets_eq_Pow by auto  | 
|
1675  | 
  { fix r let ?x = "f -` {r} \<inter> space M"
 | 
|
1676  | 
have "?x \<subseteq> space M" by auto  | 
|
1677  | 
    with finite_space sets_eq_Pow have "measure M ?x = (\<Sum>i \<in> ?x. measure M {i})"
 | 
|
1678  | 
by (auto intro!: measure_real_sum_image) }  | 
|
1679  | 
note measure_eq_setsum = this  | 
|
1680  | 
show ?thesis  | 
|
1681  | 
unfolding integral_finite[OF `f \<in> borel_measurable M` finite_space]  | 
|
1682  | 
measure_eq_setsum setsum_right_distrib  | 
|
1683  | 
apply (subst setsum_Sigma)  | 
|
1684  | 
apply (simp add: finite_space)  | 
|
1685  | 
apply (simp add: finite_space)  | 
|
1686  | 
proof (rule setsum_reindex_cong[symmetric])  | 
|
1687  | 
    fix a assume "a \<in> Sigma (f ` space M) (\<lambda>x. f -` {x} \<inter> space M)"
 | 
|
1688  | 
    thus "(\<lambda>(x, y). x * measure M {y}) a = f (snd a) * measure_space.measure M {snd a}"
 | 
|
1689  | 
by auto  | 
|
1690  | 
qed (auto intro!: image_eqI inj_onI)  | 
|
1691  | 
qed  | 
|
1692  | 
||
1693  | 
lemma (in finite_measure_space) RN_deriv_finite_singleton:  | 
|
| 35582 | 1694  | 
fixes v :: "'a set \<Rightarrow> real"  | 
| 35977 | 1695  | 
assumes ms_v: "measure_space (M\<lparr>measure := v\<rparr>)"  | 
| 36624 | 1696  | 
  and eq_0: "\<And>x. \<lbrakk> x \<in> space M ; measure M {x} = 0 \<rbrakk> \<Longrightarrow> v {x} = 0"
 | 
| 35582 | 1697  | 
  and "x \<in> space M" and "measure M {x} \<noteq> 0"
 | 
1698  | 
  shows "RN_deriv v x = v {x} / (measure M {x})" (is "_ = ?v x")
 | 
|
1699  | 
unfolding RN_deriv_def  | 
|
1700  | 
proof (rule someI2_ex[where Q = "\<lambda>f. f x = ?v x"], rule exI[where x = ?v], safe)  | 
|
1701  | 
  show "(\<lambda>a. v {a} / measure_space.measure M {a}) \<in> borel_measurable M"
 | 
|
| 35977 | 1702  | 
unfolding borel_measurable_le_iff using sets_eq_Pow by auto  | 
| 35582 | 1703  | 
next  | 
1704  | 
fix a assume "a \<in> sets M"  | 
|
1705  | 
hence "a \<subseteq> space M" and "finite a"  | 
|
| 35977 | 1706  | 
using sets_into_space finite_space by (auto intro: finite_subset)  | 
| 36624 | 1707  | 
  have *: "\<And>x a. x \<in> space M \<Longrightarrow> (if measure M {x} = 0 then 0 else v {x} * indicator_fn a x) =
 | 
| 35582 | 1708  | 
    v {x} * indicator_fn a x" using eq_0 by auto
 | 
1709  | 
||
1710  | 
from measure_space.measure_real_sum_image[OF ms_v, of a]  | 
|
| 35977 | 1711  | 
sets_eq_Pow `a \<in> sets M` sets_into_space `finite a`  | 
| 35582 | 1712  | 
  have "v a = (\<Sum>x\<in>a. v {x})" by auto
 | 
1713  | 
  thus "integral (\<lambda>x. v {x} / measure_space.measure M {x} * indicator_fn a x) = v a"
 | 
|
| 35977 | 1714  | 
apply (simp add: eq_0 integral_finite_singleton)  | 
| 35582 | 1715  | 
apply (unfold divide_1)  | 
| 35977 | 1716  | 
by (simp add: * indicator_fn_def if_distrib setsum_cases finite_space `a \<subseteq> space M` Int_absorb1)  | 
| 35582 | 1717  | 
next  | 
1718  | 
fix w assume "w \<in> borel_measurable M"  | 
|
1719  | 
assume int_eq_v: "\<forall>a\<in>sets M. integral (\<lambda>x. w x * indicator_fn a x) = v a"  | 
|
| 35977 | 1720  | 
  have "{x} \<in> sets M" using sets_eq_Pow `x \<in> space M` by auto
 | 
| 35582 | 1721  | 
|
1722  | 
  have "w x * measure M {x} =
 | 
|
1723  | 
    (\<Sum>y\<in>space M. w y * indicator_fn {x} y * measure M {y})"
 | 
|
1724  | 
apply (subst (3) mult_commute)  | 
|
| 35977 | 1725  | 
unfolding indicator_fn_def if_distrib setsum_cases[OF finite_space]  | 
| 35582 | 1726  | 
using `x \<in> space M` by simp  | 
1727  | 
  also have "... = v {x}"
 | 
|
1728  | 
    using int_eq_v[rule_format, OF `{x} \<in> sets M`]
 | 
|
| 35977 | 1729  | 
by (simp add: integral_finite_singleton)  | 
| 35582 | 1730  | 
  finally show "w x = v {x} / measure M {x}"
 | 
1731  | 
    using `measure M {x} \<noteq> 0` by (simp add: eq_divide_eq)
 | 
|
1732  | 
qed fact  | 
|
1733  | 
||
| 35748 | 1734  | 
end  |