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