| author | blanchet | 
| Thu, 02 Oct 2014 18:10:09 +0200 | |
| changeset 58517 | 64f6b4bd52a7 | 
| parent 57512 | cc97b347b301 | 
| child 58606 | 9c66f7c541fb | 
| permissions | -rw-r--r-- | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1 | (* Title: HOL/Probability/Nonnegative_Lebesgue_Integration.thy | 
| 42067 | 2 | Author: Johannes Hölzl, TU München | 
| 3 | Author: Armin Heller, TU München | |
| 4 | *) | |
| 38656 | 5 | |
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 6 | header {* Lebesgue Integration for Nonnegative Functions *}
 | 
| 35582 | 7 | |
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 8 | theory Nonnegative_Lebesgue_Integration | 
| 47694 | 9 | imports Measure_Space Borel_Space | 
| 35582 | 10 | begin | 
| 11 | ||
| 56949 | 12 | lemma indicator_less_ereal[simp]: | 
| 13 | "indicator A x \<le> (indicator B x::ereal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)" | |
| 14 | by (simp add: indicator_def not_le) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 15 | |
| 56994 | 16 | subsection "Simple function" | 
| 35582 | 17 | |
| 38656 | 18 | text {*
 | 
| 19 | ||
| 56996 | 20 | Our simple functions are not restricted to nonnegative real numbers. Instead | 
| 38656 | 21 | they are just functions with a finite range and are measurable when singleton | 
| 22 | sets are measurable. | |
| 35582 | 23 | |
| 38656 | 24 | *} | 
| 25 | ||
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 26 | definition "simple_function M g \<longleftrightarrow> | 
| 38656 | 27 | finite (g ` space M) \<and> | 
| 28 |     (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
 | |
| 36624 | 29 | |
| 47694 | 30 | lemma simple_functionD: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 31 | assumes "simple_function M g" | 
| 40875 | 32 | shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M" | 
| 40871 | 33 | proof - | 
| 34 | show "finite (g ` space M)" | |
| 35 | using assms unfolding simple_function_def by auto | |
| 40875 | 36 | have "g -` X \<inter> space M = g -` (X \<inter> g`space M) \<inter> space M" by auto | 
| 37 |   also have "\<dots> = (\<Union>x\<in>X \<inter> g`space M. g-`{x} \<inter> space M)" by auto
 | |
| 38 | finally show "g -` X \<inter> space M \<in> sets M" using assms | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 39 | by (auto simp del: UN_simps simp: simple_function_def) | 
| 40871 | 40 | qed | 
| 36624 | 41 | |
| 56949 | 42 | lemma measurable_simple_function[measurable_dest]: | 
| 43 | "simple_function M f \<Longrightarrow> f \<in> measurable M (count_space UNIV)" | |
| 44 | unfolding simple_function_def measurable_def | |
| 45 | proof safe | |
| 46 |   fix A assume "finite (f ` space M)" "\<forall>x\<in>f ` space M. f -` {x} \<inter> space M \<in> sets M"
 | |
| 47 |   then have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) \<in> sets M"
 | |
| 48 | by (intro sets.finite_UN) auto | |
| 49 |   also have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) = f -` A \<inter> space M"
 | |
| 50 | by (auto split: split_if_asm) | |
| 51 | finally show "f -` A \<inter> space M \<in> sets M" . | |
| 52 | qed simp | |
| 53 | ||
| 54 | lemma borel_measurable_simple_function: | |
| 55 | "simple_function M f \<Longrightarrow> f \<in> borel_measurable M" | |
| 56 | by (auto dest!: measurable_simple_function simp: measurable_def) | |
| 57 | ||
| 47694 | 58 | lemma simple_function_measurable2[intro]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 59 | assumes "simple_function M f" "simple_function M g" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 60 | shows "f -` A \<inter> g -` B \<inter> space M \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 61 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 62 | have "f -` A \<inter> g -` B \<inter> space M = (f -` A \<inter> space M) \<inter> (g -` B \<inter> space M)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 63 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 64 | then show ?thesis using assms[THEN simple_functionD(2)] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 65 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 66 | |
| 47694 | 67 | lemma simple_function_indicator_representation: | 
| 43920 | 68 | fixes f ::"'a \<Rightarrow> ereal" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 69 | assumes f: "simple_function M f" and x: "x \<in> space M" | 
| 38656 | 70 |   shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
 | 
| 71 | (is "?l = ?r") | |
| 72 | proof - | |
| 38705 | 73 | have "?r = (\<Sum>y \<in> f ` space M. | 
| 38656 | 74 |     (if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
 | 
| 57418 | 75 | by (auto intro!: setsum.cong) | 
| 38656 | 76 |   also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
 | 
| 57418 | 77 | using assms by (auto dest: simple_functionD simp: setsum.delta) | 
| 38656 | 78 | also have "... = f x" using x by (auto simp: indicator_def) | 
| 79 | finally show ?thesis by auto | |
| 80 | qed | |
| 36624 | 81 | |
| 47694 | 82 | lemma simple_function_notspace: | 
| 43920 | 83 | "simple_function M (\<lambda>x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h") | 
| 35692 | 84 | proof - | 
| 38656 | 85 |   have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
 | 
| 86 | hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) | |
| 87 |   have "?h -` {0} \<inter> space M = space M" by auto
 | |
| 88 | thus ?thesis unfolding simple_function_def by auto | |
| 89 | qed | |
| 90 | ||
| 47694 | 91 | lemma simple_function_cong: | 
| 38656 | 92 | assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 93 | shows "simple_function M f \<longleftrightarrow> simple_function M g" | 
| 38656 | 94 | proof - | 
| 95 | have "f ` space M = g ` space M" | |
| 96 |     "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
 | |
| 97 | using assms by (auto intro!: image_eqI) | |
| 98 | thus ?thesis unfolding simple_function_def using assms by simp | |
| 99 | qed | |
| 100 | ||
| 47694 | 101 | lemma simple_function_cong_algebra: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 102 | assumes "sets N = sets M" "space N = space M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 103 | shows "simple_function M f \<longleftrightarrow> simple_function N f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 104 | unfolding simple_function_def assms .. | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 105 | |
| 47694 | 106 | lemma simple_function_borel_measurable: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 107 |   fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
 | 
| 38656 | 108 | assumes "f \<in> borel_measurable M" and "finite (f ` space M)" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 109 | shows "simple_function M f" | 
| 38656 | 110 | using assms unfolding simple_function_def | 
| 111 | by (auto intro: borel_measurable_vimage) | |
| 112 | ||
| 56949 | 113 | lemma simple_function_eq_measurable: | 
| 43920 | 114 | fixes f :: "'a \<Rightarrow> ereal" | 
| 56949 | 115 | shows "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> measurable M (count_space UNIV)" | 
| 116 | using simple_function_borel_measurable[of f] measurable_simple_function[of M f] | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44666diff
changeset | 117 | by (fastforce simp: simple_function_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 118 | |
| 47694 | 119 | lemma simple_function_const[intro, simp]: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 120 | "simple_function M (\<lambda>x. c)" | 
| 38656 | 121 | by (auto intro: finite_subset simp: simple_function_def) | 
| 47694 | 122 | lemma simple_function_compose[intro, simp]: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 123 | assumes "simple_function M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 124 | shows "simple_function M (g \<circ> f)" | 
| 38656 | 125 | unfolding simple_function_def | 
| 126 | proof safe | |
| 127 | show "finite ((g \<circ> f) ` space M)" | |
| 56154 
f0a927235162
more complete set of lemmas wrt. image and composition
 haftmann parents: 
54611diff
changeset | 128 | using assms unfolding simple_function_def by (auto simp: image_comp [symmetric]) | 
| 38656 | 129 | next | 
| 130 | fix x assume "x \<in> space M" | |
| 131 |   let ?G = "g -` {g (f x)} \<inter> (f`space M)"
 | |
| 132 |   have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
 | |
| 133 |     (\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
 | |
| 134 |   show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
 | |
| 135 | using assms unfolding simple_function_def * | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 136 | by (rule_tac sets.finite_UN) auto | 
| 38656 | 137 | qed | 
| 138 | ||
| 47694 | 139 | lemma simple_function_indicator[intro, simp]: | 
| 38656 | 140 | assumes "A \<in> sets M" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 141 | shows "simple_function M (indicator A)" | 
| 35692 | 142 | proof - | 
| 38656 | 143 |   have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
 | 
| 144 | by (auto simp: indicator_def) | |
| 145 | hence "finite ?S" by (rule finite_subset) simp | |
| 146 | moreover have "- A \<inter> space M = space M - A" by auto | |
| 147 | ultimately show ?thesis unfolding simple_function_def | |
| 46905 | 148 | using assms by (auto simp: indicator_def [abs_def]) | 
| 35692 | 149 | qed | 
| 150 | ||
| 47694 | 151 | lemma simple_function_Pair[intro, simp]: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 152 | assumes "simple_function M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 153 | assumes "simple_function M g" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 154 | shows "simple_function M (\<lambda>x. (f x, g x))" (is "simple_function M ?p") | 
| 38656 | 155 | unfolding simple_function_def | 
| 156 | proof safe | |
| 157 | show "finite (?p ` space M)" | |
| 158 | using assms unfolding simple_function_def | |
| 159 | by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto | |
| 160 | next | |
| 161 | fix x assume "x \<in> space M" | |
| 162 |   have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
 | |
| 163 |       (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
 | |
| 164 | by auto | |
| 165 |   with `x \<in> space M` show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M"
 | |
| 166 | using assms unfolding simple_function_def by auto | |
| 167 | qed | |
| 35692 | 168 | |
| 47694 | 169 | lemma simple_function_compose1: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 170 | assumes "simple_function M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 171 | shows "simple_function M (\<lambda>x. g (f x))" | 
| 38656 | 172 | using simple_function_compose[OF assms, of g] | 
| 173 | by (simp add: comp_def) | |
| 35582 | 174 | |
| 47694 | 175 | lemma simple_function_compose2: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 176 | assumes "simple_function M f" and "simple_function M g" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 177 | shows "simple_function M (\<lambda>x. h (f x) (g x))" | 
| 38656 | 178 | proof - | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 179 | have "simple_function M ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))" | 
| 38656 | 180 | using assms by auto | 
| 181 | thus ?thesis by (simp_all add: comp_def) | |
| 182 | qed | |
| 35582 | 183 | |
| 47694 | 184 | lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"] | 
| 38656 | 185 | and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"] | 
| 186 | and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] | |
| 187 | and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"] | |
| 188 | and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"] | |
| 189 | and simple_function_inverse[intro, simp] = simple_function_compose[where g="inverse"] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 190 | and simple_function_max[intro, simp] = simple_function_compose2[where h=max] | 
| 38656 | 191 | |
| 47694 | 192 | lemma simple_function_setsum[intro, simp]: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 193 | assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 194 | shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)" | 
| 38656 | 195 | proof cases | 
| 196 | assume "finite P" from this assms show ?thesis by induct auto | |
| 197 | qed auto | |
| 35582 | 198 | |
| 56949 | 199 | lemma simple_function_ereal[intro, simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 200 | fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f" | 
| 56949 | 201 | shows "simple_function M (\<lambda>x. ereal (f x))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 202 | by (auto intro!: simple_function_compose1[OF sf]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 203 | |
| 56949 | 204 | lemma simple_function_real_of_nat[intro, simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 205 | fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f" | 
| 56949 | 206 | shows "simple_function M (\<lambda>x. real (f x))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 207 | by (auto intro!: simple_function_compose1[OF sf]) | 
| 35582 | 208 | |
| 47694 | 209 | lemma borel_measurable_implies_simple_function_sequence: | 
| 43920 | 210 | fixes u :: "'a \<Rightarrow> ereal" | 
| 38656 | 211 | assumes u: "u \<in> borel_measurable M" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 212 | shows "\<exists>f. incseq f \<and> (\<forall>i. \<infinity> \<notin> range (f i) \<and> simple_function M (f i)) \<and> | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 213 | (\<forall>x. (SUP i. f i x) = max 0 (u x)) \<and> (\<forall>i x. 0 \<le> f i x)" | 
| 35582 | 214 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 215 | def f \<equiv> "\<lambda>x i. if real i \<le> u x then i * 2 ^ i else natfloor (real (u x) * 2 ^ i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 216 |   { fix x j have "f x j \<le> j * 2 ^ j" unfolding f_def
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 217 | proof (split split_if, intro conjI impI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 218 | assume "\<not> real j \<le> u x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 219 | then have "natfloor (real (u x) * 2 ^ j) \<le> natfloor (j * 2 ^ j)" | 
| 56536 | 220 | by (cases "u x") (auto intro!: natfloor_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 221 | moreover have "real (natfloor (j * 2 ^ j)) \<le> j * 2^j" | 
| 56536 | 222 | by (intro real_natfloor_le) auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 223 | ultimately show "natfloor (real (u x) * 2 ^ j) \<le> j * 2 ^ j" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 224 | unfolding real_of_nat_le_iff by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 225 | qed auto } | 
| 38656 | 226 | note f_upper = this | 
| 35582 | 227 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 228 | have real_f: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 229 | "\<And>i x. real (f x i) = (if real i \<le> u x then i * 2 ^ i else real (natfloor (real (u x) * 2 ^ i)))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 230 | unfolding f_def by auto | 
| 35582 | 231 | |
| 46731 | 232 | let ?g = "\<lambda>j x. real (f x j) / 2^j :: ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 233 | show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 234 | proof (intro exI[of _ ?g] conjI allI ballI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 235 | fix i | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 236 | have "simple_function M (\<lambda>x. real (f x i))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 237 | proof (intro simple_function_borel_measurable) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 238 | show "(\<lambda>x. real (f x i)) \<in> borel_measurable M" | 
| 50021 
d96a3f468203
add support for function application to measurability prover
 hoelzl parents: 
50003diff
changeset | 239 | using u by (auto simp: real_f) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 240 |       have "(\<lambda>x. real (f x i))`space M \<subseteq> real`{..i*2^i}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 241 | using f_upper[of _ i] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 242 | then show "finite ((\<lambda>x. real (f x i))`space M)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 243 | by (rule finite_subset) auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 244 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 245 | then show "simple_function M (?g i)" | 
| 43920 | 246 | by (auto intro: simple_function_ereal simple_function_div) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 247 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 248 | show "incseq ?g" | 
| 43920 | 249 | proof (intro incseq_ereal incseq_SucI le_funI) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 250 | fix x and i :: nat | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 251 | have "f x i * 2 \<le> f x (Suc i)" unfolding f_def | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 252 | proof ((split split_if)+, intro conjI impI) | 
| 43920 | 253 | assume "ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 254 | then show "i * 2 ^ i * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 255 | by (cases "u x") (auto intro!: le_natfloor) | 
| 38656 | 256 | next | 
| 43920 | 257 | assume "\<not> ereal (real i) \<le> u x" "ereal (real (Suc i)) \<le> u x" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 258 | then show "natfloor (real (u x) * 2 ^ i) * 2 \<le> Suc i * 2 ^ Suc i" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 259 | by (cases "u x") auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 260 | next | 
| 43920 | 261 | assume "\<not> ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 262 | have "natfloor (real (u x) * 2 ^ i) * 2 = natfloor (real (u x) * 2 ^ i) * natfloor 2" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 263 | by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 264 | also have "\<dots> \<le> natfloor (real (u x) * 2 ^ i * 2)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 265 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 266 | assume "0 \<le> u x" then show ?thesis | 
| 46671 | 267 | by (intro le_mult_natfloor) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 268 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 269 | assume "\<not> 0 \<le> u x" then show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 270 | by (cases "u x") (auto simp: natfloor_neg mult_nonpos_nonneg) | 
| 38656 | 271 | qed | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 272 | also have "\<dots> = natfloor (real (u x) * 2 ^ Suc i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 273 | by (simp add: ac_simps) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 274 | finally show "natfloor (real (u x) * 2 ^ i) * 2 \<le> natfloor (real (u x) * 2 ^ Suc i)" . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 275 | qed simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 276 | then show "?g i x \<le> ?g (Suc i) x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 277 | by (auto simp: field_simps) | 
| 35582 | 278 | qed | 
| 38656 | 279 | next | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 280 | fix x show "(SUP i. ?g i x) = max 0 (u x)" | 
| 51000 | 281 | proof (rule SUP_eqI) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 282 | fix i show "?g i x \<le> max 0 (u x)" unfolding max_def f_def | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 283 | by (cases "u x") (auto simp: field_simps real_natfloor_le natfloor_neg | 
| 56536 | 284 | mult_nonpos_nonneg) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 285 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 286 | fix y assume *: "\<And>i. i \<in> UNIV \<Longrightarrow> ?g i x \<le> y" | 
| 56571 
f4635657d66f
added divide_nonneg_nonneg and co; made it a simp rule
 hoelzl parents: 
56537diff
changeset | 287 | have "\<And>i. 0 \<le> ?g i x" by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 288 | from order_trans[OF this *] have "0 \<le> y" by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 289 | show "max 0 (u x) \<le> y" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 290 | proof (cases y) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 291 | case (real r) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 292 | with * have *: "\<And>i. f x i \<le> r * 2^i" by (auto simp: divide_le_eq) | 
| 44666 | 293 | from reals_Archimedean2[of r] * have "u x \<noteq> \<infinity>" by (auto simp: f_def) (metis less_le_not_le) | 
| 43920 | 294 | then have "\<exists>p. max 0 (u x) = ereal p \<and> 0 \<le> p" by (cases "u x") (auto simp: max_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 295 | then guess p .. note ux = this | 
| 44666 | 296 | obtain m :: nat where m: "p < real m" using reals_Archimedean2 .. | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 297 | have "p \<le> r" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 298 | proof (rule ccontr) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 299 | assume "\<not> p \<le> r" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 300 | with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"] | 
| 56536 | 301 | obtain N where "\<forall>n\<ge>N. r * 2^n < p * 2^n - 1" by (auto simp: field_simps) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 302 | then have "r * 2^max N m < p * 2^max N m - 1" by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 303 | moreover | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 304 | have "real (natfloor (p * 2 ^ max N m)) \<le> r * 2 ^ max N m" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 305 | using *[of "max N m"] m unfolding real_f using ux | 
| 56536 | 306 | by (cases "0 \<le> u x") (simp_all add: max_def split: split_if_asm) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 307 | then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 308 | by (metis real_natfloor_gt_diff_one less_le_trans) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 309 | ultimately show False by auto | 
| 38656 | 310 | qed | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 311 | then show "max 0 (u x) \<le> y" using real ux by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 312 | qed (insert `0 \<le> y`, auto) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 313 | qed | 
| 56571 
f4635657d66f
added divide_nonneg_nonneg and co; made it a simp rule
 hoelzl parents: 
56537diff
changeset | 314 | qed auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 315 | qed | 
| 35582 | 316 | |
| 47694 | 317 | lemma borel_measurable_implies_simple_function_sequence': | 
| 43920 | 318 | fixes u :: "'a \<Rightarrow> ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 319 | assumes u: "u \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 320 | obtains f where "\<And>i. simple_function M (f i)" "incseq f" "\<And>i. \<infinity> \<notin> range (f i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 321 | "\<And>x. (SUP i. f i x) = max 0 (u x)" "\<And>i x. 0 \<le> f i x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 322 | using borel_measurable_implies_simple_function_sequence[OF u] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 323 | |
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 324 | lemma simple_function_induct[consumes 1, case_names cong set mult add, induct set: simple_function]: | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 325 | fixes u :: "'a \<Rightarrow> ereal" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 326 | assumes u: "simple_function M u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 327 | assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (AE x in M. f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 328 | assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 329 | assumes mult: "\<And>u c. P u \<Longrightarrow> P (\<lambda>x. c * u x)" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 330 | assumes add: "\<And>u v. P u \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 331 | shows "P u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 332 | proof (rule cong) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 333 |   from AE_space show "AE x in M. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x"
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 334 | proof eventually_elim | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 335 | fix x assume x: "x \<in> space M" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 336 | from simple_function_indicator_representation[OF u x] | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 337 |     show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 338 | qed | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 339 | next | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 340 | from u have "finite (u ` space M)" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 341 | unfolding simple_function_def by auto | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 342 |   then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 343 | proof induct | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 344 | case empty show ?case | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 345 |       using set[of "{}"] by (simp add: indicator_def[abs_def])
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 346 | qed (auto intro!: add mult set simple_functionD u) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 347 | next | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 348 |   show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
 | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 349 | apply (subst simple_function_cong) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 350 | apply (rule simple_function_indicator_representation[symmetric]) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 351 | apply (auto intro: u) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 352 | done | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 353 | qed fact | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 354 | |
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 355 | lemma simple_function_induct_nn[consumes 2, case_names cong set mult add]: | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 356 | fixes u :: "'a \<Rightarrow> ereal" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 357 | assumes u: "simple_function M u" and nn: "\<And>x. 0 \<le> u x" | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 358 | assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> simple_function M g \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 359 | assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)" | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 360 | assumes mult: "\<And>u c. 0 \<le> c \<Longrightarrow> simple_function M u \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 361 | assumes add: "\<And>u v. simple_function M u \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> P u \<Longrightarrow> simple_function M v \<Longrightarrow> (\<And>x. 0 \<le> v x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 362 | shows "P u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 363 | proof - | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 364 | show ?thesis | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 365 | proof (rule cong) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 366 | fix x assume x: "x \<in> space M" | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 367 | from simple_function_indicator_representation[OF u x] | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 368 |     show "(\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x) = u x" ..
 | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 369 | next | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 370 |     show "simple_function M (\<lambda>x. (\<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x))"
 | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 371 | apply (subst simple_function_cong) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 372 | apply (rule simple_function_indicator_representation[symmetric]) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 373 | apply (auto intro: u) | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 374 | done | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 375 | next | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 376 | |
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 377 | from u nn have "finite (u ` space M)" "\<And>x. x \<in> u ` space M \<Longrightarrow> 0 \<le> x" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 378 | unfolding simple_function_def by auto | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 379 |     then show "P (\<lambda>x. \<Sum>y\<in>u ` space M. y * indicator (u -` {y} \<inter> space M) x)"
 | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 380 | proof induct | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 381 | case empty show ?case | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 382 |         using set[of "{}"] by (simp add: indicator_def[abs_def])
 | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 383 | next | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 384 | case (insert x S) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 385 |       { fix z have "(\<Sum>y\<in>S. y * indicator (u -` {y} \<inter> space M) z) = 0 \<or>
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 386 |           x * indicator (u -` {x} \<inter> space M) z = 0"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 387 | using insert by (subst setsum_ereal_0) (auto simp: indicator_def) } | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 388 | note disj = this | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 389 | from insert show ?case | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 390 | by (auto intro!: add mult set simple_functionD u setsum_nonneg simple_function_setsum disj) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 391 | qed | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 392 | qed fact | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 393 | qed | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 394 | |
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 395 | lemma borel_measurable_induct[consumes 2, case_names cong set mult add seq, induct set: borel_measurable]: | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 396 | fixes u :: "'a \<Rightarrow> ereal" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 397 | assumes u: "u \<in> borel_measurable M" "\<And>x. 0 \<le> u x" | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 398 | assumes cong: "\<And>f g. f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P g \<Longrightarrow> P f" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 399 | assumes set: "\<And>A. A \<in> sets M \<Longrightarrow> P (indicator A)" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 400 | assumes mult': "\<And>u c. 0 \<le> c \<Longrightarrow> c < \<infinity> \<Longrightarrow> u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < \<infinity>) \<Longrightarrow> P u \<Longrightarrow> P (\<lambda>x. c * u x)" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 401 | assumes add: "\<And>u v. u \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> u x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x < \<infinity>) \<Longrightarrow> P u \<Longrightarrow> v \<in> borel_measurable M \<Longrightarrow> (\<And>x. 0 \<le> v x) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> v x < \<infinity>) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = 0 \<or> v x = 0) \<Longrightarrow> P v \<Longrightarrow> P (\<lambda>x. v x + u x)" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 402 | assumes seq: "\<And>U. (\<And>i. U i \<in> borel_measurable M) \<Longrightarrow> (\<And>i x. 0 \<le> U i x) \<Longrightarrow> (\<And>i x. x \<in> space M \<Longrightarrow> U i x < \<infinity>) \<Longrightarrow> (\<And>i. P (U i)) \<Longrightarrow> incseq U \<Longrightarrow> u = (SUP i. U i) \<Longrightarrow> P (SUP i. U i)" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 403 | shows "P u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 404 | using u | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 405 | proof (induct rule: borel_measurable_implies_simple_function_sequence') | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 406 | fix U assume U: "\<And>i. simple_function M (U i)" "incseq U" "\<And>i. \<infinity> \<notin> range (U i)" and | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 407 | sup: "\<And>x. (SUP i. U i x) = max 0 (u x)" and nn: "\<And>i x. 0 \<le> U i x" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 408 | have u_eq: "u = (SUP i. U i)" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 409 | using nn u sup by (auto simp: max_def) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 410 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 411 | have not_inf: "\<And>x i. x \<in> space M \<Longrightarrow> U i x < \<infinity>" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 412 | using U by (auto simp: image_iff eq_commute) | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 413 | |
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 414 | from U have "\<And>i. U i \<in> borel_measurable M" | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 415 | by (simp add: borel_measurable_simple_function) | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 416 | |
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 417 | show "P u" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 418 | unfolding u_eq | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 419 | proof (rule seq) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 420 | fix i show "P (U i)" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 421 | using `simple_function M (U i)` nn[of i] not_inf[of _ i] | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 422 | proof (induct rule: simple_function_induct_nn) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 423 | case (mult u c) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 424 | show ?case | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 425 | proof cases | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 426 |         assume "c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0)"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 427 | with mult(2) show ?thesis | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 428 |           by (intro cong[of "\<lambda>x. c * u x" "indicator {}"] set)
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 429 | (auto dest!: borel_measurable_simple_function) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 430 | next | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 431 |         assume "\<not> (c = 0 \<or> space M = {} \<or> (\<forall>x\<in>space M. u x = 0))"
 | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 432 | with mult obtain x where u_fin: "\<And>x. x \<in> space M \<Longrightarrow> u x < \<infinity>" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 433 | and x: "x \<in> space M" "u x \<noteq> 0" "c \<noteq> 0" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 434 | by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 435 | with mult have "P u" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 436 | by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 437 | from x mult(5)[OF `x \<in> space M`] mult(1) mult(3)[of x] have "c < \<infinity>" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 438 | by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 439 | with u_fin mult | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 440 | show ?thesis | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 441 | by (intro mult') (auto dest!: borel_measurable_simple_function) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 442 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 443 | qed (auto intro: cong intro!: set add dest!: borel_measurable_simple_function) | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 444 | qed fact+ | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 445 | qed | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 446 | |
| 47694 | 447 | lemma simple_function_If_set: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 448 | assumes sf: "simple_function M f" "simple_function M g" and A: "A \<inter> space M \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 449 | shows "simple_function M (\<lambda>x. if x \<in> A then f x else g x)" (is "simple_function M ?IF") | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 450 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 451 |   def F \<equiv> "\<lambda>x. f -` {x} \<inter> space M" and G \<equiv> "\<lambda>x. g -` {x} \<inter> space M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 452 | show ?thesis unfolding simple_function_def | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 453 | proof safe | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 454 | have "?IF ` space M \<subseteq> f ` space M \<union> g ` space M" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 455 | from finite_subset[OF this] assms | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 456 | show "finite (?IF ` space M)" unfolding simple_function_def by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 457 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 458 | fix x assume "x \<in> space M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 459 |     then have *: "?IF -` {?IF x} \<inter> space M = (if x \<in> A
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 460 | then ((F (f x) \<inter> (A \<inter> space M)) \<union> (G (f x) - (G (f x) \<inter> (A \<inter> space M)))) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 461 | else ((F (g x) \<inter> (A \<inter> space M)) \<union> (G (g x) - (G (g x) \<inter> (A \<inter> space M)))))" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 462 | using sets.sets_into_space[OF A] by (auto split: split_if_asm simp: G_def F_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 463 | have [intro]: "\<And>x. F x \<in> sets M" "\<And>x. G x \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 464 | unfolding F_def G_def using sf[THEN simple_functionD(2)] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 465 |     show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
 | 
| 35582 | 466 | qed | 
| 467 | qed | |
| 468 | ||
| 47694 | 469 | lemma simple_function_If: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 470 |   assumes sf: "simple_function M f" "simple_function M g" and P: "{x\<in>space M. P x} \<in> sets M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 471 | shows "simple_function M (\<lambda>x. if P x then f x else g x)" | 
| 35582 | 472 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 473 |   have "{x\<in>space M. P x} = {x. P x} \<inter> space M" by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 474 |   with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
 | 
| 38656 | 475 | qed | 
| 476 | ||
| 47694 | 477 | lemma simple_function_subalgebra: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 478 | assumes "simple_function N f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 479 | and N_subalgebra: "sets N \<subseteq> sets M" "space N = space M" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 480 | shows "simple_function M f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 481 | using assms unfolding simple_function_def by auto | 
| 39092 | 482 | |
| 47694 | 483 | lemma simple_function_comp: | 
| 484 | assumes T: "T \<in> measurable M M'" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 485 | and f: "simple_function M' f" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 486 | shows "simple_function M (\<lambda>x. f (T x))" | 
| 41661 | 487 | proof (intro simple_function_def[THEN iffD2] conjI ballI) | 
| 488 | have "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'" | |
| 489 | using T unfolding measurable_def by auto | |
| 490 | then show "finite ((\<lambda>x. f (T x)) ` space M)" | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 491 | using f unfolding simple_function_def by (auto intro: finite_subset) | 
| 41661 | 492 | fix i assume i: "i \<in> (\<lambda>x. f (T x)) ` space M" | 
| 493 | then have "i \<in> f ` space M'" | |
| 494 | using T unfolding measurable_def by auto | |
| 495 |   then have "f -` {i} \<inter> space M' \<in> sets M'"
 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 496 | using f unfolding simple_function_def by auto | 
| 41661 | 497 |   then have "T -` (f -` {i} \<inter> space M') \<inter> space M \<in> sets M"
 | 
| 498 | using T unfolding measurable_def by auto | |
| 499 |   also have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
 | |
| 500 | using T unfolding measurable_def by auto | |
| 501 |   finally show "(\<lambda>x. f (T x)) -` {i} \<inter> space M \<in> sets M" .
 | |
| 40859 | 502 | qed | 
| 503 | ||
| 56994 | 504 | subsection "Simple integral" | 
| 38656 | 505 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 506 | definition simple_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>S") where
 | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 507 |   "integral\<^sup>S M f = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M))"
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 508 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 509 | syntax | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 510 |   "_simple_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^sup>S _. _ \<partial>_" [60,61] 110)
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 511 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 512 | translations | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 513 | "\<integral>\<^sup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)" | 
| 35582 | 514 | |
| 47694 | 515 | lemma simple_integral_cong: | 
| 38656 | 516 | assumes "\<And>t. t \<in> space M \<Longrightarrow> f t = g t" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 517 | shows "integral\<^sup>S M f = integral\<^sup>S M g" | 
| 38656 | 518 | proof - | 
| 519 | have "f ` space M = g ` space M" | |
| 520 |     "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
 | |
| 521 | using assms by (auto intro!: image_eqI) | |
| 522 | thus ?thesis unfolding simple_integral_def by simp | |
| 523 | qed | |
| 524 | ||
| 47694 | 525 | lemma simple_integral_const[simp]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 526 | "(\<integral>\<^sup>Sx. c \<partial>M) = c * (emeasure M) (space M)" | 
| 38656 | 527 | proof (cases "space M = {}")
 | 
| 528 | case True thus ?thesis unfolding simple_integral_def by simp | |
| 529 | next | |
| 530 |   case False hence "(\<lambda>x. c) ` space M = {c}" by auto
 | |
| 531 | thus ?thesis unfolding simple_integral_def by simp | |
| 35582 | 532 | qed | 
| 533 | ||
| 47694 | 534 | lemma simple_function_partition: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 535 | assumes f: "simple_function M f" and g: "simple_function M g" | 
| 56949 | 536 | assumes sub: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow> g x = g y \<Longrightarrow> f x = f y" | 
| 537 | assumes v: "\<And>x. x \<in> space M \<Longrightarrow> f x = v (g x)" | |
| 538 |   shows "integral\<^sup>S M f = (\<Sum>y\<in>g ` space M. v y * emeasure M {x\<in>space M. g x = y})"
 | |
| 539 | (is "_ = ?r") | |
| 540 | proof - | |
| 541 | from f g have [simp]: "finite (f`space M)" "finite (g`space M)" | |
| 542 | by (auto simp: simple_function_def) | |
| 543 | from f g have [measurable]: "f \<in> measurable M (count_space UNIV)" "g \<in> measurable M (count_space UNIV)" | |
| 544 | by (auto intro: measurable_simple_function) | |
| 35582 | 545 | |
| 56949 | 546 |   { fix y assume "y \<in> space M"
 | 
| 547 |     then have "f ` space M \<inter> {i. \<exists>x\<in>space M. i = f x \<and> g y = g x} = {v (g y)}"
 | |
| 548 | by (auto cong: sub simp: v[symmetric]) } | |
| 549 | note eq = this | |
| 35582 | 550 | |
| 56949 | 551 | have "integral\<^sup>S M f = | 
| 552 | (\<Sum>y\<in>f`space M. y * (\<Sum>z\<in>g`space M. | |
| 553 |       if \<exists>x\<in>space M. y = f x \<and> z = g x then emeasure M {x\<in>space M. g x = z} else 0))"
 | |
| 554 | unfolding simple_integral_def | |
| 57418 | 555 | proof (safe intro!: setsum.cong ereal_left_mult_cong) | 
| 56949 | 556 | fix y assume y: "y \<in> space M" "f y \<noteq> 0" | 
| 557 |     have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} = 
 | |
| 558 |         {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
 | |
| 559 | by auto | |
| 560 |     have eq:"(\<Union>i\<in>{z. \<exists>x\<in>space M. f y = f x \<and> z = g x}. {x \<in> space M. g x = i}) =
 | |
| 561 |         f -` {f y} \<inter> space M"
 | |
| 562 | by (auto simp: eq_commute cong: sub rev_conj_cong) | |
| 563 | have "finite (g`space M)" by simp | |
| 564 |     then have "finite {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
 | |
| 565 | by (rule rev_finite_subset) auto | |
| 566 |     then show "emeasure M (f -` {f y} \<inter> space M) =
 | |
| 567 |       (\<Sum>z\<in>g ` space M. if \<exists>x\<in>space M. f y = f x \<and> z = g x then emeasure M {x \<in> space M. g x = z} else 0)"
 | |
| 57418 | 568 | apply (simp add: setsum.If_cases) | 
| 56949 | 569 | apply (subst setsum_emeasure) | 
| 570 | apply (auto simp: disjoint_family_on_def eq) | |
| 571 | done | |
| 38656 | 572 | qed | 
| 56949 | 573 | also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. | 
| 574 |       if \<exists>x\<in>space M. y = f x \<and> z = g x then y * emeasure M {x\<in>space M. g x = z} else 0))"
 | |
| 57418 | 575 | by (auto intro!: setsum.cong simp: setsum_ereal_right_distrib emeasure_nonneg) | 
| 56949 | 576 | also have "\<dots> = ?r" | 
| 57418 | 577 | by (subst setsum.commute) | 
| 578 | (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq) | |
| 56949 | 579 | finally show "integral\<^sup>S M f = ?r" . | 
| 35582 | 580 | qed | 
| 581 | ||
| 47694 | 582 | lemma simple_integral_add[simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 583 | assumes f: "simple_function M f" and "\<And>x. 0 \<le> f x" and g: "simple_function M g" and "\<And>x. 0 \<le> g x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 584 | shows "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = integral\<^sup>S M f + integral\<^sup>S M g" | 
| 35582 | 585 | proof - | 
| 56949 | 586 | have "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = | 
| 587 |     (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. (fst y + snd y) * emeasure M {x\<in>space M. (f x, g x) = y})"
 | |
| 588 | by (intro simple_function_partition) (auto intro: f g) | |
| 589 |   also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) +
 | |
| 590 |     (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y})"
 | |
| 57418 | 591 | using assms(2,4) by (auto intro!: setsum.cong ereal_left_distrib simp: setsum.distrib[symmetric]) | 
| 56949 | 592 |   also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. f x \<partial>M)"
 | 
| 593 | by (intro simple_function_partition[symmetric]) (auto intro: f g) | |
| 594 |   also have "(\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * emeasure M {x\<in>space M. (f x, g x) = y}) = (\<integral>\<^sup>Sx. g x \<partial>M)"
 | |
| 595 | by (intro simple_function_partition[symmetric]) (auto intro: f g) | |
| 596 | finally show ?thesis . | |
| 35582 | 597 | qed | 
| 598 | ||
| 47694 | 599 | lemma simple_integral_setsum[simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 600 | assumes "\<And>i x. i \<in> P \<Longrightarrow> 0 \<le> f i x" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 601 | assumes "\<And>i. i \<in> P \<Longrightarrow> simple_function M (f i)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 602 | shows "(\<integral>\<^sup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>S M (f i))" | 
| 38656 | 603 | proof cases | 
| 604 | assume "finite P" | |
| 605 | from this assms show ?thesis | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 606 | by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg) | 
| 38656 | 607 | qed auto | 
| 608 | ||
| 47694 | 609 | lemma simple_integral_mult[simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 610 | assumes f: "simple_function M f" "\<And>x. 0 \<le> f x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 611 | shows "(\<integral>\<^sup>Sx. c * f x \<partial>M) = c * integral\<^sup>S M f" | 
| 38656 | 612 | proof - | 
| 56949 | 613 |   have "(\<integral>\<^sup>Sx. c * f x \<partial>M) = (\<Sum>y\<in>f ` space M. (c * y) * emeasure M {x\<in>space M. f x = y})"
 | 
| 614 | using f by (intro simple_function_partition) auto | |
| 615 | also have "\<dots> = c * integral\<^sup>S M f" | |
| 616 | using f unfolding simple_integral_def | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57447diff
changeset | 617 | by (subst setsum_ereal_right_distrib) (auto simp: emeasure_nonneg mult.assoc Int_def conj_commute) | 
| 56949 | 618 | finally show ?thesis . | 
| 40871 | 619 | qed | 
| 620 | ||
| 47694 | 621 | lemma simple_integral_mono_AE: | 
| 56949 | 622 | assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g" | 
| 47694 | 623 | and mono: "AE x in M. f x \<le> g x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 624 | shows "integral\<^sup>S M f \<le> integral\<^sup>S M g" | 
| 40859 | 625 | proof - | 
| 56949 | 626 |   let ?\<mu> = "\<lambda>P. emeasure M {x\<in>space M. P x}"
 | 
| 627 | have "integral\<^sup>S M f = (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. fst y * ?\<mu> (\<lambda>x. (f x, g x) = y))" | |
| 628 | using f g by (intro simple_function_partition) auto | |
| 629 | also have "\<dots> \<le> (\<Sum>y\<in>(\<lambda>x. (f x, g x))`space M. snd y * ?\<mu> (\<lambda>x. (f x, g x) = y))" | |
| 630 | proof (clarsimp intro!: setsum_mono) | |
| 40859 | 631 | fix x assume "x \<in> space M" | 
| 56949 | 632 | let ?M = "?\<mu> (\<lambda>y. f y = f x \<and> g y = g x)" | 
| 633 | show "f x * ?M \<le> g x * ?M" | |
| 634 | proof cases | |
| 635 | assume "?M \<noteq> 0" | |
| 636 | then have "0 < ?M" | |
| 637 | by (simp add: less_le emeasure_nonneg) | |
| 638 | also have "\<dots> \<le> ?\<mu> (\<lambda>y. f x \<le> g x)" | |
| 639 | using mono by (intro emeasure_mono_AE) auto | |
| 640 | finally have "\<not> \<not> f x \<le> g x" | |
| 641 | by (intro notI) auto | |
| 642 | then show ?thesis | |
| 643 | by (intro ereal_mult_right_mono) auto | |
| 644 | qed simp | |
| 40859 | 645 | qed | 
| 56949 | 646 | also have "\<dots> = integral\<^sup>S M g" | 
| 647 | using f g by (intro simple_function_partition[symmetric]) auto | |
| 648 | finally show ?thesis . | |
| 40859 | 649 | qed | 
| 650 | ||
| 47694 | 651 | lemma simple_integral_mono: | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 652 | assumes "simple_function M f" and "simple_function M g" | 
| 38656 | 653 | and mono: "\<And> x. x \<in> space M \<Longrightarrow> f x \<le> g x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 654 | shows "integral\<^sup>S M f \<le> integral\<^sup>S M g" | 
| 41705 | 655 | using assms by (intro simple_integral_mono_AE) auto | 
| 35582 | 656 | |
| 47694 | 657 | lemma simple_integral_cong_AE: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 658 | assumes "simple_function M f" and "simple_function M g" | 
| 47694 | 659 | and "AE x in M. f x = g x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 660 | shows "integral\<^sup>S M f = integral\<^sup>S M g" | 
| 40859 | 661 | using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) | 
| 662 | ||
| 47694 | 663 | lemma simple_integral_cong': | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 664 | assumes sf: "simple_function M f" "simple_function M g" | 
| 47694 | 665 |   and mea: "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0"
 | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 666 | shows "integral\<^sup>S M f = integral\<^sup>S M g" | 
| 40859 | 667 | proof (intro simple_integral_cong_AE sf AE_I) | 
| 47694 | 668 |   show "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0" by fact
 | 
| 40859 | 669 |   show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
 | 
| 670 | using sf[THEN borel_measurable_simple_function] by auto | |
| 671 | qed simp | |
| 672 | ||
| 47694 | 673 | lemma simple_integral_indicator: | 
| 56949 | 674 | assumes A: "A \<in> sets M" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 675 | assumes f: "simple_function M f" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 676 | shows "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) = | 
| 56949 | 677 |     (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
 | 
| 678 | proof - | |
| 679 |   have eq: "(\<lambda>x. (f x, indicator A x)) ` space M \<inter> {x. snd x = 1} = (\<lambda>x. (f x, 1::ereal))`A"
 | |
| 680 | using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: split_if_asm) | |
| 681 |   have eq2: "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}"
 | |
| 682 | by (auto simp: image_iff) | |
| 683 | ||
| 684 | have "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) = | |
| 685 |     (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x))`space M. (fst y * snd y) * emeasure M {x\<in>space M. (f x, indicator A x) = y})"
 | |
| 686 | using assms by (intro simple_function_partition) auto | |
| 687 | also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ereal))`space M. | |
| 688 |     if snd y = 1 then fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A) else 0)"
 | |
| 57418 | 689 | by (auto simp: indicator_def split: split_if_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum.cong) | 
| 56949 | 690 |   also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, 1::ereal))`A. fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A))"
 | 
| 57418 | 691 | using assms by (subst setsum.If_cases) (auto intro!: simple_functionD(1) simp: eq) | 
| 56949 | 692 |   also have "\<dots> = (\<Sum>y\<in>fst`(\<lambda>x. (f x, 1::ereal))`A. y * emeasure M (f -` {y} \<inter> space M \<inter> A))"
 | 
| 57418 | 693 | by (subst setsum.reindex [of fst]) (auto simp: inj_on_def) | 
| 56949 | 694 |   also have "\<dots> = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
 | 
| 695 | using A[THEN sets.sets_into_space] | |
| 57418 | 696 | by (intro setsum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2) | 
| 56949 | 697 | finally show ?thesis . | 
| 38656 | 698 | qed | 
| 35582 | 699 | |
| 47694 | 700 | lemma simple_integral_indicator_only[simp]: | 
| 38656 | 701 | assumes "A \<in> sets M" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 702 | shows "integral\<^sup>S M (indicator A) = emeasure M A" | 
| 56949 | 703 | using simple_integral_indicator[OF assms, of "\<lambda>x. 1"] sets.sets_into_space[OF assms] | 
| 704 | by (simp_all add: image_constant_conv Int_absorb1 split: split_if_asm) | |
| 35582 | 705 | |
| 47694 | 706 | lemma simple_integral_null_set: | 
| 707 | assumes "simple_function M u" "\<And>x. 0 \<le> u x" and "N \<in> null_sets M" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 708 | shows "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = 0" | 
| 38656 | 709 | proof - | 
| 47694 | 710 | have "AE x in M. indicator N x = (0 :: ereal)" | 
| 711 | using `N \<in> null_sets M` by (auto simp: indicator_def intro!: AE_I[of _ _ N]) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 712 | then have "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = (\<integral>\<^sup>Sx. 0 \<partial>M)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 713 | using assms apply (intro simple_integral_cong_AE) by auto | 
| 40859 | 714 | then show ?thesis by simp | 
| 38656 | 715 | qed | 
| 35582 | 716 | |
| 47694 | 717 | lemma simple_integral_cong_AE_mult_indicator: | 
| 718 | assumes sf: "simple_function M f" and eq: "AE x in M. x \<in> S" and "S \<in> sets M" | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 719 | shows "integral\<^sup>S M f = (\<integral>\<^sup>Sx. f x * indicator S x \<partial>M)" | 
| 41705 | 720 | using assms by (intro simple_integral_cong_AE) auto | 
| 35582 | 721 | |
| 47694 | 722 | lemma simple_integral_cmult_indicator: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 723 | assumes A: "A \<in> sets M" | 
| 56949 | 724 | shows "(\<integral>\<^sup>Sx. c * indicator A x \<partial>M) = c * emeasure M A" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 725 | using simple_integral_mult[OF simple_function_indicator[OF A]] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 726 | unfolding simple_integral_indicator_only[OF A] by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 727 | |
| 56996 | 728 | lemma simple_integral_nonneg: | 
| 47694 | 729 | assumes f: "simple_function M f" and ae: "AE x in M. 0 \<le> f x" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 730 | shows "0 \<le> integral\<^sup>S M f" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 731 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 732 | have "integral\<^sup>S M (\<lambda>x. 0) \<le> integral\<^sup>S M f" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 733 | using simple_integral_mono_AE[OF _ f ae] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 734 | then show ?thesis by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 735 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 736 | |
| 56994 | 737 | subsection {* Integral on nonnegative functions *}
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 738 | |
| 56996 | 739 | definition nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>N") where
 | 
| 740 |   "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f}. integral\<^sup>S M g)"
 | |
| 35692 | 741 | |
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 742 | syntax | 
| 56996 | 743 |   "_nn_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^sup>+ _. _ \<partial>_" [60,61] 110)
 | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 744 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 745 | translations | 
| 56996 | 746 | "\<integral>\<^sup>+x. f \<partial>M" == "CONST nn_integral M (\<lambda>x. f)" | 
| 40872 | 747 | |
| 57025 | 748 | lemma nn_integral_nonneg: "0 \<le> integral\<^sup>N M f" | 
| 56996 | 749 | by (auto intro!: SUP_upper2[of "\<lambda>x. 0"] simp: nn_integral_def le_fun_def) | 
| 40873 | 750 | |
| 56996 | 751 | lemma nn_integral_not_MInfty[simp]: "integral\<^sup>N M f \<noteq> -\<infinity>" | 
| 752 | using nn_integral_nonneg[of M f] by auto | |
| 47694 | 753 | |
| 56996 | 754 | lemma nn_integral_def_finite: | 
| 755 |   "integral\<^sup>N M f = (SUP g : {g. simple_function M g \<and> g \<le> max 0 \<circ> f \<and> range g \<subseteq> {0 ..< \<infinity>}}. integral\<^sup>S M g)"
 | |
| 56218 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 haftmann parents: 
56213diff
changeset | 756 | (is "_ = SUPREMUM ?A ?f") | 
| 56996 | 757 | unfolding nn_integral_def | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 758 | proof (safe intro!: antisym SUP_least) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 759 | fix g assume g: "simple_function M g" "g \<le> max 0 \<circ> f" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 760 |   let ?G = "{x \<in> space M. \<not> g x \<noteq> \<infinity>}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 761 | note gM = g(1)[THEN borel_measurable_simple_function] | 
| 50252 | 762 | have \<mu>_G_pos: "0 \<le> (emeasure M) ?G" using gM by auto | 
| 46731 | 763 | let ?g = "\<lambda>y x. if g x = \<infinity> then y else max 0 (g x)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 764 | from g gM have g_in_A: "\<And>y. 0 \<le> y \<Longrightarrow> y \<noteq> \<infinity> \<Longrightarrow> ?g y \<in> ?A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 765 | apply (safe intro!: simple_function_max simple_function_If) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 766 | apply (force simp: max_def le_fun_def split: split_if_asm)+ | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 767 | done | 
| 56218 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 haftmann parents: 
56213diff
changeset | 768 | show "integral\<^sup>S M g \<le> SUPREMUM ?A ?f" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 769 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 770 | have g0: "?g 0 \<in> ?A" by (intro g_in_A) auto | 
| 47694 | 771 | assume "(emeasure M) ?G = 0" | 
| 772 | with gM have "AE x in M. x \<notin> ?G" | |
| 773 | by (auto simp add: AE_iff_null intro!: null_setsI) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 774 | with gM g show ?thesis | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 775 | by (intro SUP_upper2[OF g0] simple_integral_mono_AE) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 776 | (auto simp: max_def intro!: simple_function_If) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 777 | next | 
| 50252 | 778 | assume \<mu>_G: "(emeasure M) ?G \<noteq> 0" | 
| 56218 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 haftmann parents: 
56213diff
changeset | 779 | have "SUPREMUM ?A (integral\<^sup>S M) = \<infinity>" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 780 | proof (intro SUP_PInfty) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 781 | fix n :: nat | 
| 47694 | 782 | let ?y = "ereal (real n) / (if (emeasure M) ?G = \<infinity> then 1 else (emeasure M) ?G)" | 
| 50252 | 783 | have "0 \<le> ?y" "?y \<noteq> \<infinity>" using \<mu>_G \<mu>_G_pos by (auto simp: ereal_divide_eq) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 784 | then have "?g ?y \<in> ?A" by (rule g_in_A) | 
| 47694 | 785 | have "real n \<le> ?y * (emeasure M) ?G" | 
| 50252 | 786 | using \<mu>_G \<mu>_G_pos by (cases "(emeasure M) ?G") (auto simp: field_simps) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 787 | also have "\<dots> = (\<integral>\<^sup>Sx. ?y * indicator ?G x \<partial>M)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 788 | using `0 \<le> ?y` `?g ?y \<in> ?A` gM | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 789 | by (subst simple_integral_cmult_indicator) auto | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 790 | also have "\<dots> \<le> integral\<^sup>S M (?g ?y)" using `?g ?y \<in> ?A` gM | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 791 | by (intro simple_integral_mono) auto | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 792 | finally show "\<exists>i\<in>?A. real n \<le> integral\<^sup>S M i" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 793 | using `?g ?y \<in> ?A` by blast | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 794 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 795 | then show ?thesis by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 796 | qed | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 797 | qed (auto intro: SUP_upper) | 
| 40873 | 798 | |
| 56996 | 799 | lemma nn_integral_mono_AE: | 
| 800 | assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^sup>N M u \<le> integral\<^sup>N M v" | |
| 801 | unfolding nn_integral_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 802 | proof (safe intro!: SUP_mono) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 803 | fix n assume n: "simple_function M n" "n \<le> max 0 \<circ> u" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 804 | from ae[THEN AE_E] guess N . note N = this | 
| 47694 | 805 | then have ae_N: "AE x in M. x \<notin> N" by (auto intro: AE_not_in) | 
| 46731 | 806 | let ?n = "\<lambda>x. n x * indicator (space M - N) x" | 
| 47694 | 807 | have "AE x in M. n x \<le> ?n x" "simple_function M ?n" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 808 | using n N ae_N by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 809 | moreover | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 810 |   { fix x have "?n x \<le> max 0 (v x)"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 811 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 812 | assume x: "x \<in> space M - N" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 813 | with N have "u x \<le> v x" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 814 | with n(2)[THEN le_funD, of x] x show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 815 | by (auto simp: max_def split: split_if_asm) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 816 | qed simp } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 817 | then have "?n \<le> max 0 \<circ> v" by (auto simp: le_funI) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 818 | moreover have "integral\<^sup>S M n \<le> integral\<^sup>S M ?n" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 819 | using ae_N N n by (auto intro!: simple_integral_mono_AE) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 820 |   ultimately show "\<exists>m\<in>{g. simple_function M g \<and> g \<le> max 0 \<circ> v}. integral\<^sup>S M n \<le> integral\<^sup>S M m"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 821 | by force | 
| 38656 | 822 | qed | 
| 823 | ||
| 56996 | 824 | lemma nn_integral_mono: | 
| 825 | "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^sup>N M u \<le> integral\<^sup>N M v" | |
| 826 | by (auto intro: nn_integral_mono_AE) | |
| 40859 | 827 | |
| 56996 | 828 | lemma nn_integral_cong_AE: | 
| 829 | "AE x in M. u x = v x \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v" | |
| 830 | by (auto simp: eq_iff intro!: nn_integral_mono_AE) | |
| 40859 | 831 | |
| 56996 | 832 | lemma nn_integral_cong: | 
| 833 | "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v" | |
| 834 | by (auto intro: nn_integral_cong_AE) | |
| 40859 | 835 | |
| 56996 | 836 | lemma nn_integral_cong_strong: | 
| 837 | "M = N \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N N v" | |
| 838 | by (auto intro: nn_integral_cong) | |
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 839 | |
| 56996 | 840 | lemma nn_integral_eq_simple_integral: | 
| 841 | assumes f: "simple_function M f" "\<And>x. 0 \<le> f x" shows "integral\<^sup>N M f = integral\<^sup>S M f" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 842 | proof - | 
| 46731 | 843 | let ?f = "\<lambda>x. f x * indicator (space M) x" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 844 | have f': "simple_function M ?f" using f by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 845 | with f(2) have [simp]: "max 0 \<circ> ?f = ?f" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 846 | by (auto simp: fun_eq_iff max_def split: split_indicator) | 
| 56996 | 847 | have "integral\<^sup>N M ?f \<le> integral\<^sup>S M ?f" using f' | 
| 848 | by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def) | |
| 849 | moreover have "integral\<^sup>S M ?f \<le> integral\<^sup>N M ?f" | |
| 850 | unfolding nn_integral_def | |
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 851 | using f' by (auto intro!: SUP_upper) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 852 | ultimately show ?thesis | 
| 56996 | 853 | by (simp cong: nn_integral_cong simple_integral_cong) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 854 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 855 | |
| 56996 | 856 | lemma nn_integral_eq_simple_integral_AE: | 
| 857 | assumes f: "simple_function M f" "AE x in M. 0 \<le> f x" shows "integral\<^sup>N M f = integral\<^sup>S M f" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 858 | proof - | 
| 47694 | 859 | have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max) | 
| 56996 | 860 | with f have "integral\<^sup>N M f = integral\<^sup>S M (\<lambda>x. max 0 (f x))" | 
| 861 | by (simp cong: nn_integral_cong_AE simple_integral_cong_AE | |
| 862 | add: nn_integral_eq_simple_integral) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 863 | with assms show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 864 | by (auto intro!: simple_integral_cong_AE split: split_max) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 865 | qed | 
| 40873 | 866 | |
| 56996 | 867 | lemma nn_integral_SUP_approx: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 868 | assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 869 |   and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
 | 
| 56996 | 870 | shows "integral\<^sup>S M u \<le> (SUP i. integral\<^sup>N M (f i))" (is "_ \<le> ?S") | 
| 43920 | 871 | proof (rule ereal_le_mult_one_interval) | 
| 56996 | 872 | have "0 \<le> (SUP i. integral\<^sup>N M (f i))" | 
| 873 | using f(3) by (auto intro!: SUP_upper2 nn_integral_nonneg) | |
| 874 | then show "(SUP i. integral\<^sup>N M (f i)) \<noteq> -\<infinity>" by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 875 | have u_range: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> u x \<and> u x \<noteq> \<infinity>" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 876 | using u(3) by auto | 
| 43920 | 877 | fix a :: ereal assume "0 < a" "a < 1" | 
| 38656 | 878 | hence "a \<noteq> 0" by auto | 
| 46731 | 879 |   let ?B = "\<lambda>i. {x \<in> space M. a * u x \<le> f i x}"
 | 
| 38656 | 880 | have B: "\<And>i. ?B i \<in> sets M" | 
| 56949 | 881 | using f `simple_function M u`[THEN borel_measurable_simple_function] by auto | 
| 38656 | 882 | |
| 46731 | 883 | let ?uB = "\<lambda>i x. u x * indicator (?B i) x" | 
| 38656 | 884 | |
| 885 |   { fix i have "?B i \<subseteq> ?B (Suc i)"
 | |
| 886 | proof safe | |
| 887 | fix i x assume "a * u x \<le> f i x" | |
| 888 | also have "\<dots> \<le> f (Suc i) x" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 889 | using `incseq f`[THEN incseq_SucD] unfolding le_fun_def by auto | 
| 38656 | 890 | finally show "a * u x \<le> f (Suc i) x" . | 
| 891 | qed } | |
| 892 | note B_mono = this | |
| 35582 | 893 | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 894 | note B_u = sets.Int[OF u(1)[THEN simple_functionD(2)] B] | 
| 38656 | 895 | |
| 46731 | 896 |   let ?B' = "\<lambda>i n. (u -` {i} \<inter> space M) \<inter> ?B n"
 | 
| 47694 | 897 |   have measure_conv: "\<And>i. (emeasure M) (u -` {i} \<inter> space M) = (SUP n. (emeasure M) (?B' i n))"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 898 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 899 | fix i | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 900 | have 1: "range (?B' i) \<subseteq> sets M" using B_u by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 901 | have 2: "incseq (?B' i)" using B_mono by (auto intro!: incseq_SucI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 902 |     have "(\<Union>n. ?B' i n) = u -` {i} \<inter> space M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 903 | proof safe | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 904 | fix x i assume x: "x \<in> space M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 905 | show "x \<in> (\<Union>i. ?B' (u x) i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 906 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 907 | assume "u x = 0" thus ?thesis using `x \<in> space M` f(3) by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 908 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 909 | assume "u x \<noteq> 0" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 910 | with `a < 1` u_range[OF `x \<in> space M`] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 911 | have "a * u x < 1 * u x" | 
| 43920 | 912 | by (intro ereal_mult_strict_right_mono) (auto simp: image_iff) | 
| 46884 | 913 | also have "\<dots> \<le> (SUP i. f i x)" using u(2) by (auto simp: le_fun_def) | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 914 | finally obtain i where "a * u x < f i x" unfolding SUP_def | 
| 56166 | 915 | by (auto simp add: less_SUP_iff) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 916 | hence "a * u x \<le> f i x" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 917 | thus ?thesis using `x \<in> space M` by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 918 | qed | 
| 40859 | 919 | qed | 
| 47694 | 920 | then show "?thesis i" using SUP_emeasure_incseq[OF 1 2] by simp | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 921 | qed | 
| 38656 | 922 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 923 | have "integral\<^sup>S M u = (SUP i. integral\<^sup>S M (?uB i))" | 
| 41689 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 924 | unfolding simple_integral_indicator[OF B `simple_function M u`] | 
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 925 | proof (subst SUP_ereal_setsum, safe) | 
| 38656 | 926 | fix x n assume "x \<in> space M" | 
| 47694 | 927 | with u_range show "incseq (\<lambda>i. u x * (emeasure M) (?B' (u x) i))" "\<And>i. 0 \<le> u x * (emeasure M) (?B' (u x) i)" | 
| 928 | using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff) | |
| 38656 | 929 | next | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 930 | show "integral\<^sup>S M u = (\<Sum>i\<in>u ` space M. SUP n. i * (emeasure M) (?B' i n))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 931 | using measure_conv u_range B_u unfolding simple_integral_def | 
| 57418 | 932 | by (auto intro!: setsum.cong SUP_ereal_cmult [symmetric]) | 
| 38656 | 933 | qed | 
| 934 | moreover | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 935 | have "a * (SUP i. integral\<^sup>S M (?uB i)) \<le> ?S" | 
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 936 | apply (subst SUP_ereal_cmult [symmetric]) | 
| 38705 | 937 | proof (safe intro!: SUP_mono bexI) | 
| 38656 | 938 | fix i | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 939 | have "a * integral\<^sup>S M (?uB i) = (\<integral>\<^sup>Sx. a * ?uB i x \<partial>M)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 940 | using B `simple_function M u` u_range | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 941 | by (subst simple_integral_mult) (auto split: split_indicator) | 
| 56996 | 942 | also have "\<dots> \<le> integral\<^sup>N M (f i)" | 
| 38656 | 943 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 944 | have *: "simple_function M (\<lambda>x. a * ?uB i x)" using B `0 < a` u(1) by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 945 | show ?thesis using f(3) * u_range `0 < a` | 
| 56996 | 946 | by (subst nn_integral_eq_simple_integral[symmetric]) | 
| 947 | (auto intro!: nn_integral_mono split: split_indicator) | |
| 38656 | 948 | qed | 
| 56996 | 949 | finally show "a * integral\<^sup>S M (?uB i) \<le> integral\<^sup>N M (f i)" | 
| 38656 | 950 | by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 951 | next | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 952 | fix i show "0 \<le> \<integral>\<^sup>S x. ?uB i x \<partial>M" using B `0 < a` u(1) u_range | 
| 56996 | 953 | by (intro simple_integral_nonneg) (auto split: split_indicator) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 954 | qed (insert `0 < a`, auto) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 955 | ultimately show "a * integral\<^sup>S M u \<le> ?S" by simp | 
| 35582 | 956 | qed | 
| 957 | ||
| 56996 | 958 | lemma incseq_nn_integral: | 
| 959 | assumes "incseq f" shows "incseq (\<lambda>i. integral\<^sup>N M (f i))" | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 960 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 961 | have "\<And>i x. f i x \<le> f (Suc i) x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 962 | using assms by (auto dest!: incseq_SucD simp: le_fun_def) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 963 | then show ?thesis | 
| 56996 | 964 | by (auto intro!: incseq_SucI nn_integral_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 965 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 966 | |
| 35582 | 967 | text {* Beppo-Levi monotone convergence theorem *}
 | 
| 56996 | 968 | lemma nn_integral_monotone_convergence_SUP: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 969 | assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>i x. 0 \<le> f i x" | 
| 56996 | 970 | shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 971 | proof (rule antisym) | 
| 56996 | 972 | show "(SUP j. integral\<^sup>N M (f j)) \<le> (\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M)" | 
| 973 | by (auto intro!: SUP_least SUP_upper nn_integral_mono) | |
| 38656 | 974 | next | 
| 56996 | 975 | show "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP j. integral\<^sup>N M (f j))" | 
| 976 | unfolding nn_integral_def_finite[of _ "\<lambda>x. SUP i. f i x"] | |
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 977 | proof (safe intro!: SUP_least) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 978 | fix g assume g: "simple_function M g" | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 979 |       and *: "g \<le> max 0 \<circ> (\<lambda>x. SUP i. f i x)" "range g \<subseteq> {0..<\<infinity>}"
 | 
| 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 980 |     then have "\<And>x. 0 \<le> (SUP i. f i x)" and g': "g`space M \<subseteq> {0..<\<infinity>}"
 | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 981 | using f by (auto intro!: SUP_upper2) | 
| 56996 | 982 | with * show "integral\<^sup>S M g \<le> (SUP j. integral\<^sup>N M (f j))" | 
| 983 | by (intro nn_integral_SUP_approx[OF f g _ g']) | |
| 46884 | 984 | (auto simp: le_fun_def max_def) | 
| 35582 | 985 | qed | 
| 986 | qed | |
| 987 | ||
| 56996 | 988 | lemma nn_integral_monotone_convergence_SUP_AE: | 
| 47694 | 989 | assumes f: "\<And>i. AE x in M. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x" "\<And>i. f i \<in> borel_measurable M" | 
| 56996 | 990 | shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))" | 
| 40859 | 991 | proof - | 
| 47694 | 992 | from f have "AE x in M. \<forall>i. f i x \<le> f (Suc i) x \<and> 0 \<le> f i x" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 993 | by (simp add: AE_all_countable) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 994 | from this[THEN AE_E] guess N . note N = this | 
| 46731 | 995 | let ?f = "\<lambda>i x. if x \<in> space M - N then f i x else 0" | 
| 47694 | 996 | have f_eq: "AE x in M. \<forall>i. ?f i x = f i x" using N by (auto intro!: AE_I[of _ _ N]) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 997 | then have "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (SUP i. ?f i x) \<partial>M)" | 
| 56996 | 998 | by (auto intro!: nn_integral_cong_AE) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 999 | also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. ?f i x \<partial>M))" | 
| 56996 | 1000 | proof (rule nn_integral_monotone_convergence_SUP) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1001 | show "incseq ?f" using N(1) by (force intro!: incseq_SucI le_funI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1002 |     { fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1003 | using f N(3) by (intro measurable_If_set) auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1004 | fix x show "0 \<le> ?f i x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1005 | using N(1) by auto } | 
| 40859 | 1006 | qed | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1007 | also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))" | 
| 56996 | 1008 | using f_eq by (force intro!: arg_cong[where f="SUPREMUM UNIV"] nn_integral_cong_AE ext) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1009 | finally show ?thesis . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1010 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1011 | |
| 56996 | 1012 | lemma nn_integral_monotone_convergence_SUP_AE_incseq: | 
| 47694 | 1013 | assumes f: "incseq f" "\<And>i. AE x in M. 0 \<le> f i x" and borel: "\<And>i. f i \<in> borel_measurable M" | 
| 56996 | 1014 | shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1015 | using f[unfolded incseq_Suc_iff le_fun_def] | 
| 56996 | 1016 | by (intro nn_integral_monotone_convergence_SUP_AE[OF _ borel]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1017 | auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1018 | |
| 56996 | 1019 | lemma nn_integral_monotone_convergence_simple: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1020 | assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1021 | shows "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)" | 
| 56996 | 1022 | using assms unfolding nn_integral_monotone_convergence_SUP[OF f(1) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1023 | f(3)[THEN borel_measurable_simple_function] f(2)] | 
| 56996 | 1024 | by (auto intro!: nn_integral_eq_simple_integral[symmetric] arg_cong[where f="SUPREMUM UNIV"] ext) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1025 | |
| 56996 | 1026 | lemma nn_integral_max_0: | 
| 1027 | "(\<integral>\<^sup>+x. max 0 (f x) \<partial>M) = integral\<^sup>N M f" | |
| 1028 | by (simp add: le_fun_def nn_integral_def) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1029 | |
| 56996 | 1030 | lemma nn_integral_cong_pos: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1031 | assumes "\<And>x. x \<in> space M \<Longrightarrow> f x \<le> 0 \<and> g x \<le> 0 \<or> f x = g x" | 
| 56996 | 1032 | shows "integral\<^sup>N M f = integral\<^sup>N M g" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1033 | proof - | 
| 56996 | 1034 | have "integral\<^sup>N M (\<lambda>x. max 0 (f x)) = integral\<^sup>N M (\<lambda>x. max 0 (g x))" | 
| 1035 | proof (intro nn_integral_cong) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1036 | fix x assume "x \<in> space M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1037 | from assms[OF this] show "max 0 (f x) = max 0 (g x)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1038 | by (auto split: split_max) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1039 | qed | 
| 56996 | 1040 | then show ?thesis by (simp add: nn_integral_max_0) | 
| 40859 | 1041 | qed | 
| 1042 | ||
| 47694 | 1043 | lemma SUP_simple_integral_sequences: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1044 | assumes f: "incseq f" "\<And>i x. 0 \<le> f i x" "\<And>i. simple_function M (f i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1045 | and g: "incseq g" "\<And>i x. 0 \<le> g i x" "\<And>i. simple_function M (g i)" | 
| 47694 | 1046 | and eq: "AE x in M. (SUP i. f i x) = (SUP i. g i x)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1047 | shows "(SUP i. integral\<^sup>S M (f i)) = (SUP i. integral\<^sup>S M (g i))" | 
| 56218 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 haftmann parents: 
56213diff
changeset | 1048 | (is "SUPREMUM _ ?F = SUPREMUM _ ?G") | 
| 38656 | 1049 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1050 | have "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)" | 
| 56996 | 1051 | using f by (rule nn_integral_monotone_convergence_simple) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1052 | also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. g i x) \<partial>M)" | 
| 56996 | 1053 | unfolding eq[THEN nn_integral_cong_AE] .. | 
| 38656 | 1054 | also have "\<dots> = (SUP i. ?G i)" | 
| 56996 | 1055 | using g by (rule nn_integral_monotone_convergence_simple[symmetric]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1056 | finally show ?thesis by simp | 
| 38656 | 1057 | qed | 
| 1058 | ||
| 56996 | 1059 | lemma nn_integral_const[simp]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1060 | "0 \<le> c \<Longrightarrow> (\<integral>\<^sup>+ x. c \<partial>M) = c * (emeasure M) (space M)" | 
| 56996 | 1061 | by (subst nn_integral_eq_simple_integral) auto | 
| 38656 | 1062 | |
| 56996 | 1063 | lemma nn_integral_linear: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1064 | assumes f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" and "0 \<le> a" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1065 | and g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x" | 
| 56996 | 1066 | shows "(\<integral>\<^sup>+ x. a * f x + g x \<partial>M) = a * integral\<^sup>N M f + integral\<^sup>N M g" | 
| 1067 | (is "integral\<^sup>N M ?L = _") | |
| 35582 | 1068 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1069 | from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . | 
| 56996 | 1070 | note u = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1071 | from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v . | 
| 56996 | 1072 | note v = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this | 
| 46731 | 1073 | let ?L' = "\<lambda>i x. a * u i x + v i x" | 
| 38656 | 1074 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1075 | have "?L \<in> borel_measurable M" using assms by auto | 
| 38656 | 1076 | from borel_measurable_implies_simple_function_sequence'[OF this] guess l . | 
| 56996 | 1077 | note l = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1078 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1079 | have inc: "incseq (\<lambda>i. a * integral\<^sup>S M (u i))" "incseq (\<lambda>i. integral\<^sup>S M (v i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1080 | using u v `0 \<le> a` | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1081 | by (auto simp: incseq_Suc_iff le_fun_def | 
| 43920 | 1082 | intro!: add_mono ereal_mult_left_mono simple_integral_mono) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1083 | have pos: "\<And>i. 0 \<le> integral\<^sup>S M (u i)" "\<And>i. 0 \<le> integral\<^sup>S M (v i)" "\<And>i. 0 \<le> a * integral\<^sup>S M (u i)" | 
| 56996 | 1084 | using u v `0 \<le> a` by (auto simp: simple_integral_nonneg) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1085 |   { fix i from pos[of i] have "a * integral\<^sup>S M (u i) \<noteq> -\<infinity>" "integral\<^sup>S M (v i) \<noteq> -\<infinity>"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1086 | by (auto split: split_if_asm) } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1087 | note not_MInf = this | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1088 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1089 | have l': "(SUP i. integral\<^sup>S M (l i)) = (SUP i. integral\<^sup>S M (?L' i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1090 | proof (rule SUP_simple_integral_sequences[OF l(3,6,2)]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1091 | show "incseq ?L'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1092 | using u v `0 \<le> a` unfolding incseq_Suc_iff le_fun_def | 
| 56537 | 1093 | by (auto intro!: add_mono ereal_mult_left_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1094 |     { fix x
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1095 |       { fix i have "a * u i x \<noteq> -\<infinity>" "v i x \<noteq> -\<infinity>" "u i x \<noteq> -\<infinity>" using `0 \<le> a` u(6)[of i x] v(6)[of i x]
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1096 | by auto } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1097 | then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1098 | using `0 \<le> a` u(3) v(3) u(6)[of _ x] v(6)[of _ x] | 
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 1099 | by (subst SUP_ereal_cmult [symmetric, OF u(6) `0 \<le> a`]) | 
| 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 1100 | (auto intro!: SUP_ereal_add | 
| 56537 | 1101 | simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono) } | 
| 47694 | 1102 | then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1103 | unfolding l(5) using `0 \<le> a` u(5) v(5) l(5) f(2) g(2) | 
| 56537 | 1104 | by (intro AE_I2) (auto split: split_max) | 
| 38656 | 1105 | qed | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1106 | also have "\<dots> = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" | 
| 56218 
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
 haftmann parents: 
56213diff
changeset | 1107 | using u(2, 6) v(2, 6) `0 \<le> a` by (auto intro!: arg_cong[where f="SUPREMUM UNIV"] ext) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1108 | finally have "(\<integral>\<^sup>+ x. max 0 (a * f x + g x) \<partial>M) = a * (\<integral>\<^sup>+x. max 0 (f x) \<partial>M) + (\<integral>\<^sup>+x. max 0 (g x) \<partial>M)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1109 | unfolding l(5)[symmetric] u(5)[symmetric] v(5)[symmetric] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1110 | unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric] | 
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 1111 | apply (subst SUP_ereal_cmult [symmetric, OF pos(1) `0 \<le> a`]) | 
| 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 1112 | apply (subst SUP_ereal_add [symmetric, OF inc not_MInf]) . | 
| 56996 | 1113 | then show ?thesis by (simp add: nn_integral_max_0) | 
| 38656 | 1114 | qed | 
| 1115 | ||
| 56996 | 1116 | lemma nn_integral_cmult: | 
| 49775 
970964690b3d
remove some unneeded positivity assumptions; generalize some assumptions to AE; tuned proofs
 hoelzl parents: 
47761diff
changeset | 1117 | assumes f: "f \<in> borel_measurable M" "0 \<le> c" | 
| 56996 | 1118 | shows "(\<integral>\<^sup>+ x. c * f x \<partial>M) = c * integral\<^sup>N M f" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1119 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1120 | have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using `0 \<le> c` | 
| 43920 | 1121 | by (auto split: split_max simp: ereal_zero_le_0_iff) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1122 | have "(\<integral>\<^sup>+ x. c * f x \<partial>M) = (\<integral>\<^sup>+ x. c * max 0 (f x) \<partial>M)" | 
| 56996 | 1123 | by (simp add: nn_integral_max_0) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1124 | then show ?thesis | 
| 56996 | 1125 | using nn_integral_linear[OF _ _ `0 \<le> c`, of "\<lambda>x. max 0 (f x)" _ "\<lambda>x. 0"] f | 
| 1126 | by (auto simp: nn_integral_max_0) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1127 | qed | 
| 38656 | 1128 | |
| 56996 | 1129 | lemma nn_integral_multc: | 
| 49775 
970964690b3d
remove some unneeded positivity assumptions; generalize some assumptions to AE; tuned proofs
 hoelzl parents: 
47761diff
changeset | 1130 | assumes "f \<in> borel_measurable M" "0 \<le> c" | 
| 56996 | 1131 | shows "(\<integral>\<^sup>+ x. f x * c \<partial>M) = integral\<^sup>N M f * c" | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57447diff
changeset | 1132 | unfolding mult.commute[of _ c] nn_integral_cmult[OF assms] by simp | 
| 41096 | 1133 | |
| 56996 | 1134 | lemma nn_integral_indicator[simp]: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1135 | "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A x\<partial>M) = (emeasure M) A" | 
| 56996 | 1136 | by (subst nn_integral_eq_simple_integral) | 
| 49775 
970964690b3d
remove some unneeded positivity assumptions; generalize some assumptions to AE; tuned proofs
 hoelzl parents: 
47761diff
changeset | 1137 | (auto simp: simple_integral_indicator) | 
| 38656 | 1138 | |
| 56996 | 1139 | lemma nn_integral_cmult_indicator: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1140 | "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * (emeasure M) A" | 
| 56996 | 1141 | by (subst nn_integral_eq_simple_integral) | 
| 41544 | 1142 | (auto simp: simple_function_indicator simple_integral_indicator) | 
| 38656 | 1143 | |
| 56996 | 1144 | lemma nn_integral_indicator': | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1145 | assumes [measurable]: "A \<inter> space M \<in> sets M" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1146 | shows "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = emeasure M (A \<inter> space M)" | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1147 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1148 | have "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = (\<integral>\<^sup>+ x. indicator (A \<inter> space M) x \<partial>M)" | 
| 56996 | 1149 | by (intro nn_integral_cong) (simp split: split_indicator) | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1150 | also have "\<dots> = emeasure M (A \<inter> space M)" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1151 | by simp | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1152 | finally show ?thesis . | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1153 | qed | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1154 | |
| 56996 | 1155 | lemma nn_integral_add: | 
| 47694 | 1156 | assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" | 
| 1157 | and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" | |
| 56996 | 1158 | shows "(\<integral>\<^sup>+ x. f x + g x \<partial>M) = integral\<^sup>N M f + integral\<^sup>N M g" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1159 | proof - | 
| 47694 | 1160 | have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)" | 
| 56537 | 1161 | using assms by (auto split: split_max) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1162 | have "(\<integral>\<^sup>+ x. f x + g x \<partial>M) = (\<integral>\<^sup>+ x. max 0 (f x + g x) \<partial>M)" | 
| 56996 | 1163 | by (simp add: nn_integral_max_0) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1164 | also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)" | 
| 56996 | 1165 | unfolding ae[THEN nn_integral_cong_AE] .. | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1166 | also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) + (\<integral>\<^sup>+ x. max 0 (g x) \<partial>M)" | 
| 56996 | 1167 | using nn_integral_linear[of "\<lambda>x. max 0 (f x)" _ 1 "\<lambda>x. max 0 (g x)"] f g | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1168 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1169 | finally show ?thesis | 
| 56996 | 1170 | by (simp add: nn_integral_max_0) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1171 | qed | 
| 38656 | 1172 | |
| 56996 | 1173 | lemma nn_integral_setsum: | 
| 47694 | 1174 | assumes "\<And>i. i\<in>P \<Longrightarrow> f i \<in> borel_measurable M" "\<And>i. i\<in>P \<Longrightarrow> AE x in M. 0 \<le> f i x" | 
| 56996 | 1175 | shows "(\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>N M (f i))" | 
| 38656 | 1176 | proof cases | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1177 | assume f: "finite P" | 
| 47694 | 1178 | from assms have "AE x in M. \<forall>i\<in>P. 0 \<le> f i x" unfolding AE_finite_all[OF f] by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1179 | from f this assms(1) show ?thesis | 
| 38656 | 1180 | proof induct | 
| 1181 | case (insert i P) | |
| 47694 | 1182 | then have "f i \<in> borel_measurable M" "AE x in M. 0 \<le> f i x" | 
| 1183 | "(\<lambda>x. \<Sum>i\<in>P. f i x) \<in> borel_measurable M" "AE x in M. 0 \<le> (\<Sum>i\<in>P. f i x)" | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1184 | by (auto intro!: setsum_nonneg) | 
| 56996 | 1185 | from nn_integral_add[OF this] | 
| 38656 | 1186 | show ?case using insert by auto | 
| 1187 | qed simp | |
| 1188 | qed simp | |
| 1189 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1190 | lemma nn_integral_bound_simple_function: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1191 | assumes bnd: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" "\<And>x. x \<in> space M \<Longrightarrow> f x < \<infinity>" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1192 | assumes f[measurable]: "simple_function M f" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1193 |   assumes supp: "emeasure M {x\<in>space M. f x \<noteq> 0} < \<infinity>"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1194 | shows "nn_integral M f < \<infinity>" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1195 | proof cases | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1196 |   assume "space M = {}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1197 | then have "nn_integral M f = (\<integral>\<^sup>+x. 0 \<partial>M)" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1198 | by (intro nn_integral_cong) auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1199 | then show ?thesis by simp | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1200 | next | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1201 |   assume "space M \<noteq> {}"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1202 | with simple_functionD(1)[OF f] bnd have bnd: "0 \<le> Max (f`space M) \<and> Max (f`space M) < \<infinity>" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1203 | by (subst Max_less_iff) (auto simp: Max_ge_iff) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1204 | |
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1205 |   have "nn_integral M f \<le> (\<integral>\<^sup>+x. Max (f`space M) * indicator {x\<in>space M. f x \<noteq> 0} x \<partial>M)"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1206 | proof (rule nn_integral_mono) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1207 | fix x assume "x \<in> space M" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1208 |     with f show "f x \<le> Max (f ` space M) * indicator {x \<in> space M. f x \<noteq> 0} x"
 | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1209 | by (auto split: split_indicator intro!: Max_ge simple_functionD) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1210 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1211 | also have "\<dots> < \<infinity>" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1212 | using bnd supp by (subst nn_integral_cmult) auto | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1213 | finally show ?thesis . | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1214 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1215 | |
| 56996 | 1216 | lemma nn_integral_Markov_inequality: | 
| 49775 
970964690b3d
remove some unneeded positivity assumptions; generalize some assumptions to AE; tuned proofs
 hoelzl parents: 
47761diff
changeset | 1217 | assumes u: "u \<in> borel_measurable M" "AE x in M. 0 \<le> u x" and "A \<in> sets M" and c: "0 \<le> c" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1218 |   shows "(emeasure M) ({x\<in>space M. 1 \<le> c * u x} \<inter> A) \<le> c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)"
 | 
| 47694 | 1219 | (is "(emeasure M) ?A \<le> _ * ?PI") | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1220 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1221 | have "?A \<in> sets M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1222 | using `A \<in> sets M` u by auto | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1223 | hence "(emeasure M) ?A = (\<integral>\<^sup>+ x. indicator ?A x \<partial>M)" | 
| 56996 | 1224 | using nn_integral_indicator by simp | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1225 | also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)" using u c | 
| 56996 | 1226 | by (auto intro!: nn_integral_mono_AE | 
| 43920 | 1227 | simp: indicator_def ereal_zero_le_0_iff) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1228 | also have "\<dots> = c * (\<integral>\<^sup>+ x. u x * indicator A x \<partial>M)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1229 | using assms | 
| 56996 | 1230 | by (auto intro!: nn_integral_cmult simp: ereal_zero_le_0_iff) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1231 | finally show ?thesis . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1232 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1233 | |
| 56996 | 1234 | lemma nn_integral_noteq_infinite: | 
| 47694 | 1235 | assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" | 
| 56996 | 1236 | and "integral\<^sup>N M g \<noteq> \<infinity>" | 
| 47694 | 1237 | shows "AE x in M. g x \<noteq> \<infinity>" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1238 | proof (rule ccontr) | 
| 47694 | 1239 | assume c: "\<not> (AE x in M. g x \<noteq> \<infinity>)" | 
| 1240 |   have "(emeasure M) {x\<in>space M. g x = \<infinity>} \<noteq> 0"
 | |
| 1241 | using c g by (auto simp add: AE_iff_null) | |
| 1242 |   moreover have "0 \<le> (emeasure M) {x\<in>space M. g x = \<infinity>}" using g by (auto intro: measurable_sets)
 | |
| 1243 |   ultimately have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}" by auto
 | |
| 1244 |   then have "\<infinity> = \<infinity> * (emeasure M) {x\<in>space M. g x = \<infinity>}" by auto
 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1245 |   also have "\<dots> \<le> (\<integral>\<^sup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
 | 
| 56996 | 1246 | using g by (subst nn_integral_cmult_indicator) auto | 
| 1247 | also have "\<dots> \<le> integral\<^sup>N M g" | |
| 1248 | using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def) | |
| 1249 | finally show False using `integral\<^sup>N M g \<noteq> \<infinity>` by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1250 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1251 | |
| 56996 | 1252 | lemma nn_integral_PInf: | 
| 56949 | 1253 | assumes f: "f \<in> borel_measurable M" | 
| 56996 | 1254 | and not_Inf: "integral\<^sup>N M f \<noteq> \<infinity>" | 
| 56949 | 1255 |   shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
 | 
| 1256 | proof - | |
| 1257 |   have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^sup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
 | |
| 56996 | 1258 | using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets) | 
| 1259 | also have "\<dots> \<le> integral\<^sup>N M (\<lambda>x. max 0 (f x))" | |
| 1260 | by (auto intro!: nn_integral_mono simp: indicator_def max_def) | |
| 1261 |   finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^sup>N M f"
 | |
| 1262 | by (simp add: nn_integral_max_0) | |
| 56949 | 1263 |   moreover have "0 \<le> (emeasure M) (f -` {\<infinity>} \<inter> space M)"
 | 
| 1264 | by (rule emeasure_nonneg) | |
| 1265 | ultimately show ?thesis | |
| 1266 | using assms by (auto split: split_if_asm) | |
| 1267 | qed | |
| 1268 | ||
| 56996 | 1269 | lemma nn_integral_PInf_AE: | 
| 1270 | assumes "f \<in> borel_measurable M" "integral\<^sup>N M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>" | |
| 56949 | 1271 | proof (rule AE_I) | 
| 1272 |   show "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
 | |
| 56996 | 1273 | by (rule nn_integral_PInf[OF assms]) | 
| 56949 | 1274 |   show "f -` {\<infinity>} \<inter> space M \<in> sets M"
 | 
| 1275 | using assms by (auto intro: borel_measurable_vimage) | |
| 1276 | qed auto | |
| 1277 | ||
| 1278 | lemma simple_integral_PInf: | |
| 1279 | assumes "simple_function M f" "\<And>x. 0 \<le> f x" | |
| 1280 | and "integral\<^sup>S M f \<noteq> \<infinity>" | |
| 1281 |   shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
 | |
| 56996 | 1282 | proof (rule nn_integral_PInf) | 
| 56949 | 1283 | show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) | 
| 56996 | 1284 | show "integral\<^sup>N M f \<noteq> \<infinity>" | 
| 1285 | using assms by (simp add: nn_integral_eq_simple_integral) | |
| 56949 | 1286 | qed | 
| 1287 | ||
| 56996 | 1288 | lemma nn_integral_diff: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1289 | assumes f: "f \<in> borel_measurable M" | 
| 47694 | 1290 | and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" | 
| 56996 | 1291 | and fin: "integral\<^sup>N M g \<noteq> \<infinity>" | 
| 47694 | 1292 | and mono: "AE x in M. g x \<le> f x" | 
| 56996 | 1293 | shows "(\<integral>\<^sup>+ x. f x - g x \<partial>M) = integral\<^sup>N M f - integral\<^sup>N M g" | 
| 38656 | 1294 | proof - | 
| 47694 | 1295 | have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x in M. 0 \<le> f x - g x" | 
| 43920 | 1296 | using assms by (auto intro: ereal_diff_positive) | 
| 47694 | 1297 | have pos_f: "AE x in M. 0 \<le> f x" using mono g by auto | 
| 43920 | 1298 |   { fix a b :: ereal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
 | 
| 1299 | by (cases rule: ereal2_cases[of a b]) auto } | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1300 | note * = this | 
| 47694 | 1301 | then have "AE x in M. f x = f x - g x + g x" | 
| 56996 | 1302 | using mono nn_integral_noteq_infinite[OF g fin] assms by auto | 
| 1303 | then have **: "integral\<^sup>N M f = (\<integral>\<^sup>+x. f x - g x \<partial>M) + integral\<^sup>N M g" | |
| 1304 | unfolding nn_integral_add[OF diff g, symmetric] | |
| 1305 | by (rule nn_integral_cong_AE) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1306 | show ?thesis unfolding ** | 
| 56996 | 1307 | using fin nn_integral_nonneg[of M g] | 
| 1308 | by (cases rule: ereal2_cases[of "\<integral>\<^sup>+ x. f x - g x \<partial>M" "integral\<^sup>N M g"]) auto | |
| 38656 | 1309 | qed | 
| 1310 | ||
| 56996 | 1311 | lemma nn_integral_suminf: | 
| 47694 | 1312 | assumes f: "\<And>i. f i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> f i x" | 
| 56996 | 1313 | shows "(\<integral>\<^sup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>N M (f i))" | 
| 38656 | 1314 | proof - | 
| 47694 | 1315 | have all_pos: "AE x in M. \<forall>i. 0 \<le> f i x" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1316 | using assms by (auto simp: AE_all_countable) | 
| 56996 | 1317 | have "(\<Sum>i. integral\<^sup>N M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>N M (f i))" | 
| 1318 | using nn_integral_nonneg by (rule suminf_ereal_eq_SUP) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1319 | also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)" | 
| 56996 | 1320 | unfolding nn_integral_setsum[OF f] .. | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1321 | also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos | 
| 56996 | 1322 | by (intro nn_integral_monotone_convergence_SUP_AE[symmetric]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1323 | (elim AE_mp, auto simp: setsum_nonneg simp del: setsum_lessThan_Suc intro!: AE_I2 setsum_mono3) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1324 | also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos | 
| 56996 | 1325 | by (intro nn_integral_cong_AE) (auto simp: suminf_ereal_eq_SUP) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1326 | finally show ?thesis by simp | 
| 38656 | 1327 | qed | 
| 1328 | ||
| 56996 | 1329 | lemma nn_integral_mult_bounded_inf: | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1330 | assumes f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1331 | and c: "0 \<le> c" "c \<noteq> \<infinity>" and ae: "AE x in M. g x \<le> c * f x" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1332 | shows "(\<integral>\<^sup>+x. g x \<partial>M) < \<infinity>" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1333 | proof - | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1334 | have "(\<integral>\<^sup>+x. g x \<partial>M) \<le> (\<integral>\<^sup>+x. c * f x \<partial>M)" | 
| 56996 | 1335 | by (intro nn_integral_mono_AE ae) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1336 | also have "(\<integral>\<^sup>+x. c * f x \<partial>M) < \<infinity>" | 
| 56996 | 1337 | using c f by (subst nn_integral_cmult) auto | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1338 | finally show ?thesis . | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1339 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1340 | |
| 38656 | 1341 | text {* Fatou's lemma: convergence theorem on limes inferior *}
 | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1342 | |
| 56996 | 1343 | lemma nn_integral_liminf: | 
| 43920 | 1344 | fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 47694 | 1345 | assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> u i x" | 
| 56996 | 1346 | shows "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))" | 
| 38656 | 1347 | proof - | 
| 47694 | 1348 | have pos: "AE x in M. \<forall>i. 0 \<le> u i x" using u by (auto simp: AE_all_countable) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1349 | have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1350 |     (SUP n. \<integral>\<^sup>+ x. (INF i:{n..}. u i x) \<partial>M)"
 | 
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 1351 | unfolding liminf_SUP_INF using pos u | 
| 56996 | 1352 | by (intro nn_integral_monotone_convergence_SUP_AE) | 
| 44937 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 hoelzl parents: 
44928diff
changeset | 1353 | (elim AE_mp, auto intro!: AE_I2 intro: INF_greatest INF_superset_mono) | 
| 56996 | 1354 | also have "\<dots> \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))" | 
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 1355 | unfolding liminf_SUP_INF | 
| 56996 | 1356 | by (auto intro!: SUP_mono exI INF_greatest nn_integral_mono INF_lower) | 
| 38656 | 1357 | finally show ?thesis . | 
| 35582 | 1358 | qed | 
| 1359 | ||
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1360 | lemma le_Limsup: | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1361 | "F \<noteq> bot \<Longrightarrow> eventually (\<lambda>x. c \<le> g x) F \<Longrightarrow> c \<le> Limsup F g" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1362 | using Limsup_mono[of "\<lambda>_. c" g F] by (simp add: Limsup_const) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1363 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1364 | lemma Limsup_le: | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1365 | "F \<noteq> bot \<Longrightarrow> eventually (\<lambda>x. f x \<le> c) F \<Longrightarrow> Limsup F f \<le> c" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1366 | using Limsup_mono[of f "\<lambda>_. c" F] by (simp add: Limsup_const) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1367 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1368 | lemma ereal_mono_minus_cancel: | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1369 | fixes a b c :: ereal | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1370 | shows "c - a \<le> c - b \<Longrightarrow> 0 \<le> c \<Longrightarrow> c < \<infinity> \<Longrightarrow> b \<le> a" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1371 | by (cases a b c rule: ereal3_cases) auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1372 | |
| 56996 | 1373 | lemma nn_integral_limsup: | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1374 | fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1375 | assumes [measurable]: "\<And>i. u i \<in> borel_measurable M" "w \<in> borel_measurable M" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1376 | assumes bounds: "\<And>i. AE x in M. 0 \<le> u i x" "\<And>i. AE x in M. u i x \<le> w x" and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>" | 
| 56996 | 1377 | shows "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1378 | proof - | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1379 | have bnd: "AE x in M. \<forall>i. 0 \<le> u i x \<and> u i x \<le> w x" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1380 | using bounds by (auto simp: AE_all_countable) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1381 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1382 | from bounds[of 0] have w_nonneg: "AE x in M. 0 \<le> w x" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1383 | by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1384 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1385 | have "(\<integral>\<^sup>+x. w x \<partial>M) - (\<integral>\<^sup>+x. limsup (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+x. w x - limsup (\<lambda>n. u n x) \<partial>M)" | 
| 56996 | 1386 | proof (intro nn_integral_diff[symmetric]) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1387 | show "AE x in M. 0 \<le> limsup (\<lambda>n. u n x)" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1388 | using bnd by (auto intro!: le_Limsup) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1389 | show "AE x in M. limsup (\<lambda>n. u n x) \<le> w x" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1390 | using bnd by (auto intro!: Limsup_le) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1391 | then have "(\<integral>\<^sup>+x. limsup (\<lambda>n. u n x) \<partial>M) < \<infinity>" | 
| 56996 | 1392 | by (intro nn_integral_mult_bounded_inf[OF _ w, of 1]) auto | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1393 | then show "(\<integral>\<^sup>+x. limsup (\<lambda>n. u n x) \<partial>M) \<noteq> \<infinity>" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1394 | by simp | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1395 | qed auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1396 | also have "\<dots> = (\<integral>\<^sup>+x. liminf (\<lambda>n. w x - u n x) \<partial>M)" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1397 | using w_nonneg | 
| 56996 | 1398 | by (intro nn_integral_cong_AE, eventually_elim) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1399 | (auto intro!: liminf_ereal_cminus[symmetric]) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1400 | also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^sup>+x. w x - u n x \<partial>M)" | 
| 56996 | 1401 | proof (rule nn_integral_liminf) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1402 | fix i show "AE x in M. 0 \<le> w x - u i x" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1403 | using bounds[of i] by eventually_elim (auto intro: ereal_diff_positive) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1404 | qed simp | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1405 | also have "(\<lambda>n. \<integral>\<^sup>+x. w x - u n x \<partial>M) = (\<lambda>n. (\<integral>\<^sup>+x. w x \<partial>M) - (\<integral>\<^sup>+x. u n x \<partial>M))" | 
| 56996 | 1406 | proof (intro ext nn_integral_diff) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1407 | fix i have "(\<integral>\<^sup>+x. u i x \<partial>M) < \<infinity>" | 
| 56996 | 1408 | using bounds by (intro nn_integral_mult_bounded_inf[OF _ w, of 1]) auto | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1409 | then show "(\<integral>\<^sup>+x. u i x \<partial>M) \<noteq> \<infinity>" by simp | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1410 | qed (insert bounds, auto) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1411 | also have "liminf (\<lambda>n. (\<integral>\<^sup>+x. w x \<partial>M) - (\<integral>\<^sup>+x. u n x \<partial>M)) = (\<integral>\<^sup>+x. w x \<partial>M) - limsup (\<lambda>n. \<integral>\<^sup>+x. u n x \<partial>M)" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1412 | using w by (intro liminf_ereal_cminus) auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1413 | finally show ?thesis | 
| 56996 | 1414 | by (rule ereal_mono_minus_cancel) (intro w nn_integral_nonneg)+ | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1415 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1416 | |
| 57025 | 1417 | lemma nn_integral_LIMSEQ: | 
| 1418 | assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>n x. 0 \<le> f n x" | |
| 1419 | and u: "\<And>x. (\<lambda>i. f i x) ----> u x" | |
| 1420 | shows "(\<lambda>n. integral\<^sup>N M (f n)) ----> integral\<^sup>N M u" | |
| 1421 | proof - | |
| 1422 | have "(\<lambda>n. integral\<^sup>N M (f n)) ----> (SUP n. integral\<^sup>N M (f n))" | |
| 1423 | using f by (intro LIMSEQ_SUP[of "\<lambda>n. integral\<^sup>N M (f n)"] incseq_nn_integral) | |
| 1424 | also have "(SUP n. integral\<^sup>N M (f n)) = integral\<^sup>N M (\<lambda>x. SUP n. f n x)" | |
| 1425 | using f by (intro nn_integral_monotone_convergence_SUP[symmetric]) | |
| 1426 | also have "integral\<^sup>N M (\<lambda>x. SUP n. f n x) = integral\<^sup>N M (\<lambda>x. u x)" | |
| 1427 | using f by (subst SUP_Lim_ereal[OF _ u]) (auto simp: incseq_def le_fun_def) | |
| 1428 | finally show ?thesis . | |
| 1429 | qed | |
| 1430 | ||
| 56996 | 1431 | lemma nn_integral_dominated_convergence: | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1432 | assumes [measurable]: | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1433 | "\<And>i. u i \<in> borel_measurable M" "u' \<in> borel_measurable M" "w \<in> borel_measurable M" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1434 | and bound: "\<And>j. AE x in M. 0 \<le> u j x" "\<And>j. AE x in M. u j x \<le> w x" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1435 | and w: "(\<integral>\<^sup>+x. w x \<partial>M) < \<infinity>" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1436 | and u': "AE x in M. (\<lambda>i. u i x) ----> u' x" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1437 | shows "(\<lambda>i. (\<integral>\<^sup>+x. u i x \<partial>M)) ----> (\<integral>\<^sup>+x. u' x \<partial>M)" | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1438 | proof - | 
| 56996 | 1439 | have "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)" | 
| 1440 | by (intro nn_integral_limsup[OF _ _ bound w]) auto | |
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1441 | moreover have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)" | 
| 56996 | 1442 | using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1443 | moreover have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)" | 
| 56996 | 1444 | using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) | 
| 1445 | moreover have "(\<integral>\<^sup>+x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))" | |
| 1446 | by (intro nn_integral_liminf[OF _ bound(1)]) auto | |
| 1447 | moreover have "liminf (\<lambda>n. integral\<^sup>N M (u n)) \<le> limsup (\<lambda>n. integral\<^sup>N M (u n))" | |
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1448 | by (intro Liminf_le_Limsup sequentially_bot) | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1449 | ultimately show ?thesis | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1450 | by (intro Liminf_eq_Limsup) auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1451 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1452 | |
| 56996 | 1453 | lemma nn_integral_null_set: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1454 | assumes "N \<in> null_sets M" shows "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = 0" | 
| 38656 | 1455 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1456 | have "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)" | 
| 56996 | 1457 | proof (intro nn_integral_cong_AE AE_I) | 
| 40859 | 1458 |     show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
 | 
| 1459 | by (auto simp: indicator_def) | |
| 47694 | 1460 | show "(emeasure M) N = 0" "N \<in> sets M" | 
| 40859 | 1461 | using assms by auto | 
| 35582 | 1462 | qed | 
| 40859 | 1463 | then show ?thesis by simp | 
| 38656 | 1464 | qed | 
| 35582 | 1465 | |
| 56996 | 1466 | lemma nn_integral_0_iff: | 
| 47694 | 1467 | assumes u: "u \<in> borel_measurable M" and pos: "AE x in M. 0 \<le> u x" | 
| 56996 | 1468 |   shows "integral\<^sup>N M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
 | 
| 47694 | 1469 | (is "_ \<longleftrightarrow> (emeasure M) ?A = 0") | 
| 35582 | 1470 | proof - | 
| 56996 | 1471 | have u_eq: "(\<integral>\<^sup>+ x. u x * indicator ?A x \<partial>M) = integral\<^sup>N M u" | 
| 1472 | by (auto intro!: nn_integral_cong simp: indicator_def) | |
| 38656 | 1473 | show ?thesis | 
| 1474 | proof | |
| 47694 | 1475 | assume "(emeasure M) ?A = 0" | 
| 56996 | 1476 | with nn_integral_null_set[of ?A M u] u | 
| 1477 | show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def) | |
| 38656 | 1478 | next | 
| 43920 | 1479 |     { fix r :: ereal and n :: nat assume gt_1: "1 \<le> real n * r"
 | 
| 1480 | then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def) | |
| 1481 | then have "0 \<le> r" by (auto simp add: ereal_zero_less_0_iff) } | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1482 | note gt_1 = this | 
| 56996 | 1483 | assume *: "integral\<^sup>N M u = 0" | 
| 46731 | 1484 |     let ?M = "\<lambda>n. {x \<in> space M. 1 \<le> real (n::nat) * u x}"
 | 
| 47694 | 1485 | have "0 = (SUP n. (emeasure M) (?M n \<inter> ?A))" | 
| 38656 | 1486 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1487 |       { fix n :: nat
 | 
| 56996 | 1488 | from nn_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"] | 
| 47694 | 1489 | have "(emeasure M) (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp | 
| 1490 | moreover have "0 \<le> (emeasure M) (?M n \<inter> ?A)" using u by auto | |
| 1491 | ultimately have "(emeasure M) (?M n \<inter> ?A) = 0" by auto } | |
| 38656 | 1492 | thus ?thesis by simp | 
| 35582 | 1493 | qed | 
| 47694 | 1494 | also have "\<dots> = (emeasure M) (\<Union>n. ?M n \<inter> ?A)" | 
| 1495 | proof (safe intro!: SUP_emeasure_incseq) | |
| 38656 | 1496 | fix n show "?M n \<inter> ?A \<in> sets M" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 1497 | using u by (auto intro!: sets.Int) | 
| 38656 | 1498 | next | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1499 |       show "incseq (\<lambda>n. {x \<in> space M. 1 \<le> real n * u x} \<inter> {x \<in> space M. u x \<noteq> 0})"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1500 | proof (safe intro!: incseq_SucI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1501 | fix n :: nat and x | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1502 | assume *: "1 \<le> real n * u x" | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 1503 | also from gt_1[OF *] have "real n * u x \<le> real (Suc n) * u x" | 
| 43920 | 1504 | using `0 \<le> u x` by (auto intro!: ereal_mult_right_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1505 | finally show "1 \<le> real (Suc n) * u x" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1506 | qed | 
| 38656 | 1507 | qed | 
| 47694 | 1508 |     also have "\<dots> = (emeasure M) {x\<in>space M. 0 < u x}"
 | 
| 1509 | proof (safe intro!: arg_cong[where f="(emeasure M)"] dest!: gt_1) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1510 | fix x assume "0 < u x" and [simp, intro]: "x \<in> space M" | 
| 38656 | 1511 | show "x \<in> (\<Union>n. ?M n \<inter> ?A)" | 
| 1512 | proof (cases "u x") | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1513 | case (real r) with `0 < u x` have "0 < r" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1514 | obtain j :: nat where "1 / r \<le> real j" using real_arch_simple .. | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1515 | hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using `0 < r` by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1516 | hence "1 \<le> real j * r" using real `0 < r` by auto | 
| 43920 | 1517 | thus ?thesis using `0 < r` real by (auto simp: one_ereal_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1518 | qed (insert `0 < u x`, auto) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1519 | qed auto | 
| 47694 | 1520 |     finally have "(emeasure M) {x\<in>space M. 0 < u x} = 0" by simp
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1521 | moreover | 
| 47694 | 1522 | from pos have "AE x in M. \<not> (u x < 0)" by auto | 
| 1523 |     then have "(emeasure M) {x\<in>space M. u x < 0} = 0"
 | |
| 1524 | using AE_iff_null[of M] u by auto | |
| 1525 |     moreover have "(emeasure M) {x\<in>space M. u x \<noteq> 0} = (emeasure M) {x\<in>space M. u x < 0} + (emeasure M) {x\<in>space M. 0 < u x}"
 | |
| 1526 | using u by (subst plus_emeasure) (auto intro!: arg_cong[where f="emeasure M"]) | |
| 1527 | ultimately show "(emeasure M) ?A = 0" by simp | |
| 35582 | 1528 | qed | 
| 1529 | qed | |
| 1530 | ||
| 56996 | 1531 | lemma nn_integral_0_iff_AE: | 
| 41705 | 1532 | assumes u: "u \<in> borel_measurable M" | 
| 56996 | 1533 | shows "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x \<le> 0)" | 
| 41705 | 1534 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1535 |   have sets: "{x\<in>space M. max 0 (u x) \<noteq> 0} \<in> sets M"
 | 
| 41705 | 1536 | using u by auto | 
| 56996 | 1537 | from nn_integral_0_iff[of "\<lambda>x. max 0 (u x)"] | 
| 1538 | have "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. max 0 (u x) = 0)" | |
| 1539 | unfolding nn_integral_max_0 | |
| 47694 | 1540 | using AE_iff_null[OF sets] u by auto | 
| 1541 | also have "\<dots> \<longleftrightarrow> (AE x in M. u x \<le> 0)" by (auto split: split_max) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1542 | finally show ?thesis . | 
| 41705 | 1543 | qed | 
| 1544 | ||
| 56996 | 1545 | lemma AE_iff_nn_integral: | 
| 1546 |   "{x\<in>space M. P x} \<in> sets M \<Longrightarrow> (AE x in M. P x) \<longleftrightarrow> integral\<^sup>N M (indicator {x. \<not> P x}) = 0"
 | |
| 1547 | by (subst nn_integral_0_iff_AE) (auto simp: one_ereal_def zero_ereal_def | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 1548 | sets.sets_Collect_neg indicator_def[abs_def] measurable_If) | 
| 50001 
382bd3173584
add syntax and a.e.-rules for (conditional) probability on predicates
 hoelzl parents: 
49800diff
changeset | 1549 | |
| 56996 | 1550 | lemma nn_integral_const_If: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1551 | "(\<integral>\<^sup>+x. a \<partial>M) = (if 0 \<le> a then a * (emeasure M) (space M) else 0)" | 
| 56996 | 1552 | by (auto intro!: nn_integral_0_iff_AE[THEN iffD2]) | 
| 42991 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1553 | |
| 56996 | 1554 | lemma nn_integral_subalgebra: | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1555 | assumes f: "f \<in> borel_measurable N" "\<And>x. 0 \<le> f x" | 
| 47694 | 1556 | and N: "sets N \<subseteq> sets M" "space N = space M" "\<And>A. A \<in> sets N \<Longrightarrow> emeasure N A = emeasure M A" | 
| 56996 | 1557 | shows "integral\<^sup>N N f = integral\<^sup>N M f" | 
| 39092 | 1558 | proof - | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1559 | have [simp]: "\<And>f :: 'a \<Rightarrow> ereal. f \<in> borel_measurable N \<Longrightarrow> f \<in> borel_measurable M" | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1560 | using N by (auto simp: measurable_def) | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1561 | have [simp]: "\<And>P. (AE x in N. P x) \<Longrightarrow> (AE x in M. P x)" | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1562 | using N by (auto simp add: eventually_ae_filter null_sets_def) | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1563 | have [simp]: "\<And>A. A \<in> sets N \<Longrightarrow> A \<in> sets M" | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1564 | using N by auto | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1565 | from f show ?thesis | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1566 | apply induct | 
| 56996 | 1567 | apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N) | 
| 1568 | apply (auto intro!: nn_integral_cong cong: nn_integral_cong simp: N(2)[symmetric]) | |
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1569 | done | 
| 39092 | 1570 | qed | 
| 1571 | ||
| 56996 | 1572 | lemma nn_integral_nat_function: | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1573 | fixes f :: "'a \<Rightarrow> nat" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1574 | assumes "f \<in> measurable M (count_space UNIV)" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1575 |   shows "(\<integral>\<^sup>+x. ereal (of_nat (f x)) \<partial>M) = (\<Sum>t. emeasure M {x\<in>space M. t < f x})"
 | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1576 | proof - | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1577 |   def F \<equiv> "\<lambda>i. {x\<in>space M. i < f x}"
 | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1578 | with assms have [measurable]: "\<And>i. F i \<in> sets M" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1579 | by auto | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1580 | |
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1581 |   { fix x assume "x \<in> space M"
 | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1582 | have "(\<lambda>i. if i < f x then 1 else 0) sums (of_nat (f x)::real)" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1583 | using sums_If_finite[of "\<lambda>i. i < f x" "\<lambda>_. 1::real"] by simp | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1584 | then have "(\<lambda>i. ereal(if i < f x then 1 else 0)) sums (ereal(of_nat(f x)))" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1585 | unfolding sums_ereal . | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1586 | moreover have "\<And>i. ereal (if i < f x then 1 else 0) = indicator (F i) x" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1587 | using `x \<in> space M` by (simp add: one_ereal_def F_def) | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1588 | ultimately have "ereal(of_nat(f x)) = (\<Sum>i. indicator (F i) x)" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1589 | by (simp add: sums_iff) } | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1590 | then have "(\<integral>\<^sup>+x. ereal (of_nat (f x)) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)" | 
| 56996 | 1591 | by (simp cong: nn_integral_cong) | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1592 | also have "\<dots> = (\<Sum>i. emeasure M (F i))" | 
| 56996 | 1593 | by (simp add: nn_integral_suminf) | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1594 | finally show ?thesis | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1595 | by (simp add: F_def) | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1596 | qed | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1597 | |
| 56994 | 1598 | subsection {* Integral under concrete measures *}
 | 
| 1599 | ||
| 1600 | subsubsection {* Distributions *}
 | |
| 47694 | 1601 | |
| 56996 | 1602 | lemma nn_integral_distr': | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1603 | assumes T: "T \<in> measurable M M'" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1604 | and f: "f \<in> borel_measurable (distr M M' T)" "\<And>x. 0 \<le> f x" | 
| 56996 | 1605 | shows "integral\<^sup>N (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)" | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1606 | using f | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1607 | proof induct | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1608 | case (cong f g) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1609 | with T show ?case | 
| 56996 | 1610 | apply (subst nn_integral_cong[of _ f g]) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1611 | apply simp | 
| 56996 | 1612 | apply (subst nn_integral_cong[of _ "\<lambda>x. f (T x)" "\<lambda>x. g (T x)"]) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1613 | apply (simp add: measurable_def Pi_iff) | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1614 | apply simp | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1615 | done | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1616 | next | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1617 | case (set A) | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1618 | then have eq: "\<And>x. x \<in> space M \<Longrightarrow> indicator A (T x) = indicator (T -` A \<inter> space M) x" | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1619 | by (auto simp: indicator_def) | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1620 | from set T show ?case | 
| 56996 | 1621 | by (subst nn_integral_cong[OF eq]) | 
| 1622 | (auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets) | |
| 1623 | qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add | |
| 1624 | nn_integral_monotone_convergence_SUP le_fun_def incseq_def) | |
| 47694 | 1625 | |
| 56996 | 1626 | lemma nn_integral_distr: | 
| 1627 | "T \<in> measurable M M' \<Longrightarrow> f \<in> borel_measurable M' \<Longrightarrow> integral\<^sup>N (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)" | |
| 1628 | by (subst (1 2) nn_integral_max_0[symmetric]) | |
| 1629 | (simp add: nn_integral_distr') | |
| 35692 | 1630 | |
| 56994 | 1631 | subsubsection {* Counting space *}
 | 
| 47694 | 1632 | |
| 1633 | lemma simple_function_count_space[simp]: | |
| 1634 | "simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)" | |
| 1635 | unfolding simple_function_def by simp | |
| 1636 | ||
| 56996 | 1637 | lemma nn_integral_count_space: | 
| 47694 | 1638 |   assumes A: "finite {a\<in>A. 0 < f a}"
 | 
| 56996 | 1639 | shows "integral\<^sup>N (count_space A) f = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)" | 
| 35582 | 1640 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1641 | have *: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>count_space A) = | 
| 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1642 |     (\<integral>\<^sup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
 | 
| 56996 | 1643 | by (auto intro!: nn_integral_cong | 
| 57418 | 1644 | simp add: indicator_def if_distrib setsum.If_cases[OF A] max_def le_less) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1645 |   also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. \<integral>\<^sup>+ x. f a * indicator {a} x \<partial>count_space A)"
 | 
| 56996 | 1646 | by (subst nn_integral_setsum) | 
| 47694 | 1647 | (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le) | 
| 1648 | also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)" | |
| 57418 | 1649 | by (auto intro!: setsum.cong simp: nn_integral_cmult_indicator one_ereal_def[symmetric]) | 
| 56996 | 1650 | finally show ?thesis by (simp add: nn_integral_max_0) | 
| 47694 | 1651 | qed | 
| 1652 | ||
| 56996 | 1653 | lemma nn_integral_count_space_finite: | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1654 | "finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. max 0 (f a))" | 
| 56996 | 1655 | by (subst nn_integral_max_0[symmetric]) | 
| 57418 | 1656 | (auto intro!: setsum.mono_neutral_left simp: nn_integral_count_space less_le) | 
| 47694 | 1657 | |
| 54418 | 1658 | lemma emeasure_UN_countable: | 
| 1659 | assumes sets: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets M" and I: "countable I" | |
| 1660 | assumes disj: "disjoint_family_on X I" | |
| 1661 | shows "emeasure M (UNION I X) = (\<integral>\<^sup>+i. emeasure M (X i) \<partial>count_space I)" | |
| 1662 | proof cases | |
| 1663 | assume "finite I" with sets disj show ?thesis | |
| 1664 | by (subst setsum_emeasure[symmetric]) | |
| 57418 | 1665 | (auto intro!: setsum.cong simp add: max_def subset_eq nn_integral_count_space_finite emeasure_nonneg) | 
| 54418 | 1666 | next | 
| 1667 | assume f: "\<not> finite I" | |
| 1668 |   then have [intro]: "I \<noteq> {}" by auto
 | |
| 1669 | from from_nat_into_inj_infinite[OF I f] from_nat_into[OF this] disj | |
| 1670 | have disj2: "disjoint_family (\<lambda>i. X (from_nat_into I i))" | |
| 1671 | unfolding disjoint_family_on_def by metis | |
| 1672 | ||
| 1673 | from f have "bij_betw (from_nat_into I) UNIV I" | |
| 1674 | using bij_betw_from_nat_into[OF I] by simp | |
| 1675 | then have "(\<Union>i\<in>I. X i) = (\<Union>i. (X \<circ> from_nat_into I) i)" | |
| 56154 
f0a927235162
more complete set of lemmas wrt. image and composition
 haftmann parents: 
54611diff
changeset | 1676 | unfolding SUP_def image_comp [symmetric] by (simp add: bij_betw_def) | 
| 54418 | 1677 | then have "emeasure M (UNION I X) = emeasure M (\<Union>i. X (from_nat_into I i))" | 
| 1678 | by simp | |
| 1679 | also have "\<dots> = (\<Sum>i. emeasure M (X (from_nat_into I i)))" | |
| 1680 |     by (intro suminf_emeasure[symmetric] disj disj2) (auto intro!: sets from_nat_into[OF `I \<noteq> {}`])
 | |
| 1681 |   also have "\<dots> = (\<Sum>n. \<integral>\<^sup>+i. emeasure M (X i) * indicator {from_nat_into I n} i \<partial>count_space I)"
 | |
| 1682 | proof (intro arg_cong[where f=suminf] ext) | |
| 1683 | fix i | |
| 1684 |     have eq: "{a \<in> I. 0 < emeasure M (X a) * indicator {from_nat_into I i} a}
 | |
| 1685 |      = (if 0 < emeasure M (X (from_nat_into I i)) then {from_nat_into I i} else {})"
 | |
| 1686 | using ereal_0_less_1 | |
| 1687 |      by (auto simp: ereal_zero_less_0_iff indicator_def from_nat_into `I \<noteq> {}` simp del: ereal_0_less_1)
 | |
| 1688 |     have "(\<integral>\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \<partial>count_space I) =
 | |
| 1689 | (if 0 < emeasure M (X (from_nat_into I i)) then emeasure M (X (from_nat_into I i)) else 0)" | |
| 56996 | 1690 | by (subst nn_integral_count_space) (simp_all add: eq) | 
| 54418 | 1691 | also have "\<dots> = emeasure M (X (from_nat_into I i))" | 
| 1692 | by (simp add: less_le emeasure_nonneg) | |
| 1693 | finally show "emeasure M (X (from_nat_into I i)) = | |
| 1694 |          \<integral>\<^sup>+ ia. emeasure M (X ia) * indicator {from_nat_into I i} ia \<partial>count_space I" ..
 | |
| 1695 | qed | |
| 1696 | also have "\<dots> = (\<integral>\<^sup>+i. emeasure M (X i) \<partial>count_space I)" | |
| 56996 | 1697 | apply (subst nn_integral_suminf[symmetric]) | 
| 1698 | apply (auto simp: emeasure_nonneg intro!: nn_integral_cong) | |
| 54418 | 1699 | proof - | 
| 1700 | fix x assume "x \<in> I" | |
| 1701 |     then have "(\<Sum>i. emeasure M (X x) * indicator {from_nat_into I i} x) = (\<Sum>i\<in>{to_nat_on I x}. emeasure M (X x) * indicator {from_nat_into I i} x)"
 | |
| 1702 | by (intro suminf_finite) (auto simp: indicator_def I f) | |
| 1703 | also have "\<dots> = emeasure M (X x)" | |
| 1704 | by (simp add: I f `x\<in>I`) | |
| 1705 |     finally show "(\<Sum>i. emeasure M (X x) * indicator {from_nat_into I i} x) = emeasure M (X x)" .
 | |
| 1706 | qed | |
| 1707 | finally show ?thesis . | |
| 1708 | qed | |
| 1709 | ||
| 57025 | 1710 | lemma nn_integral_count_space_nat: | 
| 1711 | fixes f :: "nat \<Rightarrow> ereal" | |
| 1712 | assumes nonneg: "\<And>i. 0 \<le> f i" | |
| 1713 | shows "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = (\<Sum>i. f i)" | |
| 1714 | proof - | |
| 1715 | have "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = | |
| 1716 |     (\<integral>\<^sup>+i. (\<Sum>j. f j * indicator {j} i) \<partial>count_space UNIV)"
 | |
| 1717 | proof (intro nn_integral_cong) | |
| 1718 | fix i | |
| 1719 |     have "f i = (\<Sum>j\<in>{i}. f j * indicator {j} i)"
 | |
| 1720 | by simp | |
| 1721 |     also have "\<dots> = (\<Sum>j. f j * indicator {j} i)"
 | |
| 1722 | by (rule suminf_finite[symmetric]) auto | |
| 1723 |     finally show "f i = (\<Sum>j. f j * indicator {j} i)" .
 | |
| 1724 | qed | |
| 1725 |   also have "\<dots> = (\<Sum>j. (\<integral>\<^sup>+i. f j * indicator {j} i \<partial>count_space UNIV))"
 | |
| 1726 | by (rule nn_integral_suminf) (auto simp: nonneg) | |
| 1727 | also have "\<dots> = (\<Sum>j. f j)" | |
| 1728 | by (simp add: nonneg nn_integral_cmult_indicator one_ereal_def[symmetric]) | |
| 1729 | finally show ?thesis . | |
| 1730 | qed | |
| 1731 | ||
| 1732 | lemma emeasure_countable_singleton: | |
| 1733 |   assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" and X: "countable X"
 | |
| 1734 |   shows "emeasure M X = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
 | |
| 1735 | proof - | |
| 1736 |   have "emeasure M (\<Union>i\<in>X. {i}) = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
 | |
| 1737 | using assms by (intro emeasure_UN_countable) (auto simp: disjoint_family_on_def) | |
| 1738 |   also have "(\<Union>i\<in>X. {i}) = X" by auto
 | |
| 1739 | finally show ?thesis . | |
| 1740 | qed | |
| 1741 | ||
| 1742 | lemma measure_eqI_countable: | |
| 1743 | assumes [simp]: "sets M = Pow A" "sets N = Pow A" and A: "countable A" | |
| 1744 |   assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
 | |
| 1745 | shows "M = N" | |
| 1746 | proof (rule measure_eqI) | |
| 1747 | fix X assume "X \<in> sets M" | |
| 1748 | then have X: "X \<subseteq> A" by auto | |
| 1749 | moreover with A have "countable X" by (auto dest: countable_subset) | |
| 1750 | ultimately have | |
| 1751 |     "emeasure M X = (\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X)"
 | |
| 1752 |     "emeasure N X = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
 | |
| 1753 | by (auto intro!: emeasure_countable_singleton) | |
| 1754 |   moreover have "(\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X) = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
 | |
| 1755 | using X by (intro nn_integral_cong eq) auto | |
| 1756 | ultimately show "emeasure M X = emeasure N X" | |
| 1757 | by simp | |
| 1758 | qed simp | |
| 1759 | ||
| 56994 | 1760 | subsubsection {* Measures with Restricted Space *}
 | 
| 54417 | 1761 | |
| 57137 | 1762 | lemma simple_function_iff_borel_measurable: | 
| 1763 |   fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
 | |
| 1764 | shows "simple_function M f \<longleftrightarrow> finite (f ` space M) \<and> f \<in> borel_measurable M" | |
| 1765 | by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable) | |
| 1766 | ||
| 1767 | lemma simple_function_restrict_space_ereal: | |
| 1768 | fixes f :: "'a \<Rightarrow> ereal" | |
| 1769 | assumes "\<Omega> \<inter> space M \<in> sets M" | |
| 1770 | shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 1771 | proof - | |
| 1772 |   { assume "finite (f ` space (restrict_space M \<Omega>))"
 | |
| 1773 |     then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
 | |
| 1774 | then have "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)" | |
| 1775 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 1776 | moreover | |
| 1777 |   { assume "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)"
 | |
| 1778 | then have "finite (f ` space (restrict_space M \<Omega>))" | |
| 1779 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 1780 | ultimately show ?thesis | |
| 1781 | unfolding simple_function_iff_borel_measurable | |
| 1782 | borel_measurable_restrict_space_iff_ereal[OF assms] | |
| 1783 | by auto | |
| 1784 | qed | |
| 1785 | ||
| 1786 | lemma simple_function_restrict_space: | |
| 1787 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | |
| 1788 | assumes "\<Omega> \<inter> space M \<in> sets M" | |
| 1789 | shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" | |
| 1790 | proof - | |
| 1791 |   { assume "finite (f ` space (restrict_space M \<Omega>))"
 | |
| 1792 |     then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
 | |
| 1793 | then have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)" | |
| 1794 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 1795 | moreover | |
| 1796 |   { assume "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)"
 | |
| 1797 | then have "finite (f ` space (restrict_space M \<Omega>))" | |
| 1798 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 1799 | ultimately show ?thesis | |
| 1800 | unfolding simple_function_iff_borel_measurable | |
| 1801 | borel_measurable_restrict_space_iff[OF assms] | |
| 1802 | by auto | |
| 1803 | qed | |
| 1804 | ||
| 1805 | ||
| 1806 | lemma split_indicator_asm: "P (indicator Q x) = (\<not> (x \<in> Q \<and> \<not> P 1 \<or> x \<notin> Q \<and> \<not> P 0))" | |
| 1807 | by (auto split: split_indicator) | |
| 1808 | ||
| 1809 | lemma simple_integral_restrict_space: | |
| 1810 | assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" "simple_function (restrict_space M \<Omega>) f" | |
| 1811 | shows "simple_integral (restrict_space M \<Omega>) f = simple_integral M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 1812 | using simple_function_restrict_space_ereal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)] | |
| 1813 | by (auto simp add: space_restrict_space emeasure_restrict_space[OF \<Omega>(1)] le_infI2 simple_integral_def | |
| 1814 | split: split_indicator split_indicator_asm | |
| 57418 | 1815 | intro!: setsum.mono_neutral_cong_left ereal_left_mult_cong arg_cong2[where f=emeasure]) | 
| 57137 | 1816 | |
| 1817 | lemma one_not_less_zero[simp]: "\<not> 1 < (0::ereal)" | |
| 1818 | by (simp add: zero_ereal_def one_ereal_def) | |
| 1819 | ||
| 56996 | 1820 | lemma nn_integral_restrict_space: | 
| 57137 | 1821 | assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M" | 
| 1822 | shows "nn_integral (restrict_space M \<Omega>) f = nn_integral M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 1823 | proof - | |
| 1824 |   let ?R = "restrict_space M \<Omega>" and ?X = "\<lambda>M f. {s. simple_function M s \<and> s \<le> max 0 \<circ> f \<and> range s \<subseteq> {0 ..< \<infinity>}}"
 | |
| 1825 | have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 1826 | proof (safe intro!: image_eqI) | |
| 1827 |     fix s assume s: "simple_function ?R s" "s \<le> max 0 \<circ> f" "range s \<subseteq> {0..<\<infinity>}"
 | |
| 1828 | from s show "integral\<^sup>S (restrict_space M \<Omega>) s = integral\<^sup>S M (\<lambda>x. s x * indicator \<Omega> x)" | |
| 1829 | by (intro simple_integral_restrict_space) auto | |
| 1830 | from s show "simple_function M (\<lambda>x. s x * indicator \<Omega> x)" | |
| 1831 | by (simp add: simple_function_restrict_space_ereal) | |
| 1832 | from s show "(\<lambda>x. s x * indicator \<Omega> x) \<le> max 0 \<circ> (\<lambda>x. f x * indicator \<Omega> x)" | |
| 1833 |       "\<And>x. s x * indicator \<Omega> x \<in> {0..<\<infinity>}"
 | |
| 1834 | by (auto split: split_indicator simp: le_fun_def image_subset_iff) | |
| 1835 | next | |
| 1836 |     fix s assume s: "simple_function M s" "s \<le> max 0 \<circ> (\<lambda>x. f x * indicator \<Omega> x)" "range s \<subseteq> {0..<\<infinity>}"
 | |
| 1837 | then have "simple_function M (\<lambda>x. s x * indicator (\<Omega> \<inter> space M) x)" (is ?s') | |
| 1838 | by (intro simple_function_mult simple_function_indicator) auto | |
| 1839 | also have "?s' \<longleftrightarrow> simple_function M (\<lambda>x. s x * indicator \<Omega> x)" | |
| 1840 | by (rule simple_function_cong) (auto split: split_indicator) | |
| 1841 | finally show sf: "simple_function (restrict_space M \<Omega>) s" | |
| 1842 | by (simp add: simple_function_restrict_space_ereal) | |
| 1843 | ||
| 1844 | from s have s_eq: "s = (\<lambda>x. s x * indicator \<Omega> x)" | |
| 1845 | by (auto simp add: fun_eq_iff le_fun_def image_subset_iff | |
| 1846 | split: split_indicator split_indicator_asm | |
| 1847 | intro: antisym) | |
| 1848 | ||
| 1849 | show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \<Omega>) s" | |
| 1850 | by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \<Omega> sf]) | |
| 1851 |     show "\<And>x. s x \<in> {0..<\<infinity>}"
 | |
| 1852 | using s by (auto simp: image_subset_iff) | |
| 1853 | from s show "s \<le> max 0 \<circ> f" | |
| 1854 | by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm) | |
| 1855 | qed | |
| 1856 | then show ?thesis | |
| 1857 | unfolding nn_integral_def_finite SUP_def by simp | |
| 54417 | 1858 | qed | 
| 1859 | ||
| 56994 | 1860 | subsubsection {* Measure spaces with an associated density *}
 | 
| 47694 | 1861 | |
| 1862 | definition density :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1863 | "density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)" | 
| 35582 | 1864 | |
| 47694 | 1865 | lemma | 
| 1866 | shows sets_density[simp]: "sets (density M f) = sets M" | |
| 1867 | and space_density[simp]: "space (density M f) = space M" | |
| 1868 | by (auto simp: density_def) | |
| 1869 | ||
| 50003 | 1870 | (* FIXME: add conversion to simplify space, sets and measurable *) | 
| 1871 | lemma space_density_imp[measurable_dest]: | |
| 1872 | "\<And>x M f. x \<in> space (density M f) \<Longrightarrow> x \<in> space M" by auto | |
| 1873 | ||
| 47694 | 1874 | lemma | 
| 1875 | shows measurable_density_eq1[simp]: "g \<in> measurable (density Mg f) Mg' \<longleftrightarrow> g \<in> measurable Mg Mg'" | |
| 1876 | and measurable_density_eq2[simp]: "h \<in> measurable Mh (density Mh' f) \<longleftrightarrow> h \<in> measurable Mh Mh'" | |
| 1877 | and simple_function_density_eq[simp]: "simple_function (density Mu f) u \<longleftrightarrow> simple_function Mu u" | |
| 1878 | unfolding measurable_def simple_function_def by simp_all | |
| 1879 | ||
| 1880 | lemma density_cong: "f \<in> borel_measurable M \<Longrightarrow> f' \<in> borel_measurable M \<Longrightarrow> | |
| 1881 | (AE x in M. f x = f' x) \<Longrightarrow> density M f = density M f'" | |
| 56996 | 1882 | unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed) | 
| 47694 | 1883 | |
| 1884 | lemma density_max_0: "density M f = density M (\<lambda>x. max 0 (f x))" | |
| 1885 | proof - | |
| 1886 | have "\<And>x A. max 0 (f x) * indicator A x = max 0 (f x * indicator A x)" | |
| 1887 | by (auto simp: indicator_def) | |
| 1888 | then show ?thesis | |
| 56996 | 1889 | unfolding density_def by (simp add: nn_integral_max_0) | 
| 47694 | 1890 | qed | 
| 1891 | ||
| 1892 | lemma density_ereal_max_0: "density M (\<lambda>x. ereal (f x)) = density M (\<lambda>x. ereal (max 0 (f x)))" | |
| 1893 | by (subst density_max_0) (auto intro!: arg_cong[where f="density M"] split: split_max) | |
| 38656 | 1894 | |
| 47694 | 1895 | lemma emeasure_density: | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1896 | assumes f[measurable]: "f \<in> borel_measurable M" and A[measurable]: "A \<in> sets M" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1897 | shows "emeasure (density M f) A = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M)" | 
| 47694 | 1898 | (is "_ = ?\<mu> A") | 
| 1899 | unfolding density_def | |
| 1900 | proof (rule emeasure_measure_of_sigma) | |
| 1901 | show "sigma_algebra (space M) (sets M)" .. | |
| 1902 | show "positive (sets M) ?\<mu>" | |
| 56996 | 1903 | using f by (auto simp: positive_def intro!: nn_integral_nonneg) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1904 | have \<mu>_eq: "?\<mu> = (\<lambda>A. \<integral>\<^sup>+ x. max 0 (f x) * indicator A x \<partial>M)" (is "?\<mu> = ?\<mu>'") | 
| 56996 | 1905 | apply (subst nn_integral_max_0[symmetric]) | 
| 1906 | apply (intro ext nn_integral_cong_AE AE_I2) | |
| 47694 | 1907 | apply (auto simp: indicator_def) | 
| 1908 | done | |
| 1909 | show "countably_additive (sets M) ?\<mu>" | |
| 1910 | unfolding \<mu>_eq | |
| 1911 | proof (intro countably_additiveI) | |
| 1912 | fix A :: "nat \<Rightarrow> 'a set" assume "range A \<subseteq> sets M" | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1913 | then have "\<And>i. A i \<in> sets M" by auto | 
| 47694 | 1914 | then have *: "\<And>i. (\<lambda>x. max 0 (f x) * indicator (A i) x) \<in> borel_measurable M" | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1915 | by (auto simp: set_eq_iff) | 
| 47694 | 1916 | assume disj: "disjoint_family A" | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1917 | have "(\<Sum>n. ?\<mu>' (A n)) = (\<integral>\<^sup>+ x. (\<Sum>n. max 0 (f x) * indicator (A n) x) \<partial>M)" | 
| 56996 | 1918 | using f * by (simp add: nn_integral_suminf) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1919 | also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) * (\<Sum>n. indicator (A n) x) \<partial>M)" using f | 
| 56996 | 1920 | by (auto intro!: suminf_cmult_ereal nn_integral_cong_AE) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1921 | also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) * indicator (\<Union>n. A n) x \<partial>M)" | 
| 47694 | 1922 | unfolding suminf_indicator[OF disj] .. | 
| 1923 | finally show "(\<Sum>n. ?\<mu>' (A n)) = ?\<mu>' (\<Union>x. A x)" by simp | |
| 1924 | qed | |
| 1925 | qed fact | |
| 38656 | 1926 | |
| 47694 | 1927 | lemma null_sets_density_iff: | 
| 1928 | assumes f: "f \<in> borel_measurable M" | |
| 1929 | shows "A \<in> null_sets (density M f) \<longleftrightarrow> A \<in> sets M \<and> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" | |
| 1930 | proof - | |
| 1931 |   { assume "A \<in> sets M"
 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1932 | have eq: "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. max 0 (f x) * indicator A x \<partial>M)" | 
| 56996 | 1933 | apply (subst nn_integral_max_0[symmetric]) | 
| 1934 | apply (intro nn_integral_cong) | |
| 47694 | 1935 | apply (auto simp: indicator_def) | 
| 1936 | done | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1937 | have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> | 
| 47694 | 1938 |       emeasure M {x \<in> space M. max 0 (f x) * indicator A x \<noteq> 0} = 0"
 | 
| 1939 | unfolding eq | |
| 1940 | using f `A \<in> sets M` | |
| 56996 | 1941 | by (intro nn_integral_0_iff) auto | 
| 47694 | 1942 | also have "\<dots> \<longleftrightarrow> (AE x in M. max 0 (f x) * indicator A x = 0)" | 
| 1943 | using f `A \<in> sets M` | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 1944 | by (intro AE_iff_measurable[OF _ refl, symmetric]) auto | 
| 47694 | 1945 | also have "(AE x in M. max 0 (f x) * indicator A x = 0) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" | 
| 1946 | by (auto simp add: indicator_def max_def split: split_if_asm) | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1947 | finally have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> f x \<le> 0)" . } | 
| 47694 | 1948 | with f show ?thesis | 
| 1949 | by (simp add: null_sets_def emeasure_density cong: conj_cong) | |
| 1950 | qed | |
| 1951 | ||
| 1952 | lemma AE_density: | |
| 1953 | assumes f: "f \<in> borel_measurable M" | |
| 1954 | shows "(AE x in density M f. P x) \<longleftrightarrow> (AE x in M. 0 < f x \<longrightarrow> P x)" | |
| 1955 | proof | |
| 1956 | assume "AE x in density M f. P x" | |
| 1957 |   with f obtain N where "{x \<in> space M. \<not> P x} \<subseteq> N" "N \<in> sets M" and ae: "AE x in M. x \<in> N \<longrightarrow> f x \<le> 0"
 | |
| 1958 | by (auto simp: eventually_ae_filter null_sets_density_iff) | |
| 1959 | then have "AE x in M. x \<notin> N \<longrightarrow> P x" by auto | |
| 1960 | with ae show "AE x in M. 0 < f x \<longrightarrow> P x" | |
| 1961 | by (rule eventually_elim2) auto | |
| 1962 | next | |
| 1963 | fix N assume ae: "AE x in M. 0 < f x \<longrightarrow> P x" | |
| 1964 |   then obtain N where "{x \<in> space M. \<not> (0 < f x \<longrightarrow> P x)} \<subseteq> N" "N \<in> null_sets M"
 | |
| 1965 | by (auto simp: eventually_ae_filter) | |
| 1966 |   then have *: "{x \<in> space (density M f). \<not> P x} \<subseteq> N \<union> {x\<in>space M. \<not> 0 < f x}"
 | |
| 1967 |     "N \<union> {x\<in>space M. \<not> 0 < f x} \<in> sets M" and ae2: "AE x in M. x \<notin> N"
 | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 1968 | using f by (auto simp: subset_eq intro!: sets.sets_Collect_neg AE_not_in) | 
| 47694 | 1969 | show "AE x in density M f. P x" | 
| 1970 | using ae2 | |
| 1971 | unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] | |
| 1972 |     by (intro exI[of _ "N \<union> {x\<in>space M. \<not> 0 < f x}"] conjI *)
 | |
| 1973 | (auto elim: eventually_elim2) | |
| 35582 | 1974 | qed | 
| 1975 | ||
| 56996 | 1976 | lemma nn_integral_density': | 
| 47694 | 1977 | assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1978 | assumes g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x" | 
| 56996 | 1979 | shows "integral\<^sup>N (density M f) g = (\<integral>\<^sup>+ x. f x * g x \<partial>M)" | 
| 49798 | 1980 | using g proof induct | 
| 1981 | case (cong u v) | |
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1982 | then show ?case | 
| 56996 | 1983 | apply (subst nn_integral_cong[OF cong(3)]) | 
| 1984 | apply (simp_all cong: nn_integral_cong) | |
| 49798 | 1985 | done | 
| 1986 | next | |
| 1987 | case (set A) then show ?case | |
| 1988 | by (simp add: emeasure_density f) | |
| 1989 | next | |
| 1990 | case (mult u c) | |
| 1991 | moreover have "\<And>x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) | |
| 1992 | ultimately show ?case | |
| 56996 | 1993 | using f by (simp add: nn_integral_cmult) | 
| 49798 | 1994 | next | 
| 1995 | case (add u v) | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 1996 | then have "\<And>x. f x * (v x + u x) = f x * v x + f x * u x" | 
| 49798 | 1997 | by (simp add: ereal_right_distrib) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 1998 | with add f show ?case | 
| 56996 | 1999 | by (auto simp add: nn_integral_add ereal_zero_le_0_iff intro!: nn_integral_add[symmetric]) | 
| 49798 | 2000 | next | 
| 2001 | case (seq U) | |
| 2002 | from f(2) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" | |
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 2003 | by eventually_elim (simp add: SUP_ereal_cmult seq) | 
| 49798 | 2004 | from seq f show ?case | 
| 56996 | 2005 | apply (simp add: nn_integral_monotone_convergence_SUP) | 
| 2006 | apply (subst nn_integral_cong_AE[OF eq]) | |
| 2007 | apply (subst nn_integral_monotone_convergence_SUP_AE) | |
| 49798 | 2008 | apply (auto simp: incseq_def le_fun_def intro!: ereal_mult_left_mono) | 
| 2009 | done | |
| 47694 | 2010 | qed | 
| 38705 | 2011 | |
| 56996 | 2012 | lemma nn_integral_density: | 
| 49798 | 2013 | "f \<in> borel_measurable M \<Longrightarrow> AE x in M. 0 \<le> f x \<Longrightarrow> g' \<in> borel_measurable M \<Longrightarrow> | 
| 56996 | 2014 | integral\<^sup>N (density M f) g' = (\<integral>\<^sup>+ x. f x * g' x \<partial>M)" | 
| 2015 | by (subst (1 2) nn_integral_max_0[symmetric]) | |
| 2016 | (auto intro!: nn_integral_cong_AE | |
| 2017 | simp: measurable_If max_def ereal_zero_le_0_iff nn_integral_density') | |
| 49798 | 2018 | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2019 | lemma density_distr: | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2020 | assumes [measurable]: "f \<in> borel_measurable N" "X \<in> measurable M N" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2021 | shows "density (distr M N X) f = distr (density M (\<lambda>x. f (X x))) N X" | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2022 | by (intro measure_eqI) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2023 | (auto simp add: emeasure_density nn_integral_distr emeasure_distr | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2024 | split: split_indicator intro!: nn_integral_cong) | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2025 | |
| 47694 | 2026 | lemma emeasure_restricted: | 
| 2027 | assumes S: "S \<in> sets M" and X: "X \<in> sets M" | |
| 2028 | shows "emeasure (density M (indicator S)) X = emeasure M (S \<inter> X)" | |
| 38705 | 2029 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2030 | have "emeasure (density M (indicator S)) X = (\<integral>\<^sup>+x. indicator S x * indicator X x \<partial>M)" | 
| 47694 | 2031 | using S X by (simp add: emeasure_density) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2032 | also have "\<dots> = (\<integral>\<^sup>+x. indicator (S \<inter> X) x \<partial>M)" | 
| 56996 | 2033 | by (auto intro!: nn_integral_cong simp: indicator_def) | 
| 47694 | 2034 | also have "\<dots> = emeasure M (S \<inter> X)" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 2035 | using S X by (simp add: sets.Int) | 
| 47694 | 2036 | finally show ?thesis . | 
| 2037 | qed | |
| 2038 | ||
| 2039 | lemma measure_restricted: | |
| 2040 | "S \<in> sets M \<Longrightarrow> X \<in> sets M \<Longrightarrow> measure (density M (indicator S)) X = measure M (S \<inter> X)" | |
| 2041 | by (simp add: emeasure_restricted measure_def) | |
| 2042 | ||
| 2043 | lemma (in finite_measure) finite_measure_restricted: | |
| 2044 | "S \<in> sets M \<Longrightarrow> finite_measure (density M (indicator S))" | |
| 2045 | by default (simp add: emeasure_restricted) | |
| 2046 | ||
| 2047 | lemma emeasure_density_const: | |
| 2048 | "A \<in> sets M \<Longrightarrow> 0 \<le> c \<Longrightarrow> emeasure (density M (\<lambda>_. c)) A = c * emeasure M A" | |
| 56996 | 2049 | by (auto simp: nn_integral_cmult_indicator emeasure_density) | 
| 47694 | 2050 | |
| 2051 | lemma measure_density_const: | |
| 2052 | "A \<in> sets M \<Longrightarrow> 0 < c \<Longrightarrow> c \<noteq> \<infinity> \<Longrightarrow> measure (density M (\<lambda>_. c)) A = real c * measure M A" | |
| 2053 | by (auto simp: emeasure_density_const measure_def) | |
| 2054 | ||
| 2055 | lemma density_density_eq: | |
| 2056 | "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. 0 \<le> f x \<Longrightarrow> | |
| 2057 | density (density M f) g = density M (\<lambda>x. f x * g x)" | |
| 56996 | 2058 | by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps) | 
| 47694 | 2059 | |
| 2060 | lemma distr_density_distr: | |
| 2061 | assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M" | |
| 2062 | and inv: "\<forall>x\<in>space M. T' (T x) = x" | |
| 2063 | assumes f: "f \<in> borel_measurable M'" | |
| 2064 | shows "distr (density (distr M M' T) f) M T' = density M (f \<circ> T)" (is "?R = ?L") | |
| 2065 | proof (rule measure_eqI) | |
| 2066 | fix A assume A: "A \<in> sets ?R" | |
| 2067 |   { fix x assume "x \<in> space M"
 | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 2068 | with sets.sets_into_space[OF A] | 
| 47694 | 2069 | have "indicator (T' -` A \<inter> space M') (T x) = (indicator A x :: ereal)" | 
| 2070 | using T inv by (auto simp: indicator_def measurable_space) } | |
| 2071 | with A T T' f show "emeasure ?R A = emeasure ?L A" | |
| 2072 | by (simp add: measurable_comp emeasure_density emeasure_distr | |
| 56996 | 2073 | nn_integral_distr measurable_sets cong: nn_integral_cong) | 
| 47694 | 2074 | qed simp | 
| 2075 | ||
| 2076 | lemma density_density_divide: | |
| 2077 | fixes f g :: "'a \<Rightarrow> real" | |
| 2078 | assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" | |
| 2079 | assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" | |
| 2080 | assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0" | |
| 2081 | shows "density (density M f) (\<lambda>x. g x / f x) = density M g" | |
| 2082 | proof - | |
| 2083 | have "density M g = density M (\<lambda>x. f x * (g x / f x))" | |
| 2084 | using f g ac by (auto intro!: density_cong measurable_If) | |
| 2085 | then show ?thesis | |
| 2086 | using f g by (subst density_density_eq) auto | |
| 38705 | 2087 | qed | 
| 2088 | ||
| 56994 | 2089 | subsubsection {* Point measure *}
 | 
| 47694 | 2090 | |
| 2091 | definition point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
 | |
| 2092 | "point_measure A f = density (count_space A) f" | |
| 2093 | ||
| 2094 | lemma | |
| 2095 | shows space_point_measure: "space (point_measure A f) = A" | |
| 2096 | and sets_point_measure: "sets (point_measure A f) = Pow A" | |
| 2097 | by (auto simp: point_measure_def) | |
| 2098 | ||
| 2099 | lemma measurable_point_measure_eq1[simp]: | |
| 2100 | "g \<in> measurable (point_measure A f) M \<longleftrightarrow> g \<in> A \<rightarrow> space M" | |
| 2101 | unfolding point_measure_def by simp | |
| 2102 | ||
| 2103 | lemma measurable_point_measure_eq2_finite[simp]: | |
| 2104 | "finite A \<Longrightarrow> | |
| 2105 | g \<in> measurable M (point_measure A f) \<longleftrightarrow> | |
| 2106 |     (g \<in> space M \<rightarrow> A \<and> (\<forall>a\<in>A. g -` {a} \<inter> space M \<in> sets M))"
 | |
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 2107 | unfolding point_measure_def by (simp add: measurable_count_space_eq2) | 
| 47694 | 2108 | |
| 2109 | lemma simple_function_point_measure[simp]: | |
| 2110 | "simple_function (point_measure A f) g \<longleftrightarrow> finite (g ` A)" | |
| 2111 | by (simp add: point_measure_def) | |
| 2112 | ||
| 2113 | lemma emeasure_point_measure: | |
| 2114 |   assumes A: "finite {a\<in>X. 0 < f a}" "X \<subseteq> A"
 | |
| 2115 | shows "emeasure (point_measure A f) X = (\<Sum>a|a\<in>X \<and> 0 < f a. f a)" | |
| 35977 | 2116 | proof - | 
| 47694 | 2117 |   have "{a. (a \<in> X \<longrightarrow> a \<in> A \<and> 0 < f a) \<and> a \<in> X} = {a\<in>X. 0 < f a}"
 | 
| 2118 | using `X \<subseteq> A` by auto | |
| 2119 | with A show ?thesis | |
| 56996 | 2120 | by (simp add: emeasure_density nn_integral_count_space ereal_zero_le_0_iff | 
| 47694 | 2121 | point_measure_def indicator_def) | 
| 35977 | 2122 | qed | 
| 2123 | ||
| 47694 | 2124 | lemma emeasure_point_measure_finite: | 
| 49795 | 2125 | "finite A \<Longrightarrow> (\<And>i. i \<in> A \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)" | 
| 57418 | 2126 | by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le) | 
| 47694 | 2127 | |
| 49795 | 2128 | lemma emeasure_point_measure_finite2: | 
| 2129 | "X \<subseteq> A \<Longrightarrow> finite X \<Longrightarrow> (\<And>i. i \<in> X \<Longrightarrow> 0 \<le> f i) \<Longrightarrow> emeasure (point_measure A f) X = (\<Sum>a\<in>X. f a)" | |
| 2130 | by (subst emeasure_point_measure) | |
| 57418 | 2131 | (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le) | 
| 49795 | 2132 | |
| 47694 | 2133 | lemma null_sets_point_measure_iff: | 
| 2134 | "X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x \<le> 0)" | |
| 2135 | by (auto simp: AE_count_space null_sets_density_iff point_measure_def) | |
| 2136 | ||
| 2137 | lemma AE_point_measure: | |
| 2138 | "(AE x in point_measure A f. P x) \<longleftrightarrow> (\<forall>x\<in>A. 0 < f x \<longrightarrow> P x)" | |
| 2139 | unfolding point_measure_def | |
| 2140 | by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) | |
| 2141 | ||
| 56996 | 2142 | lemma nn_integral_point_measure: | 
| 47694 | 2143 |   "finite {a\<in>A. 0 < f a \<and> 0 < g a} \<Longrightarrow>
 | 
| 56996 | 2144 | integral\<^sup>N (point_measure A f) g = (\<Sum>a|a\<in>A \<and> 0 < f a \<and> 0 < g a. f a * g a)" | 
| 47694 | 2145 | unfolding point_measure_def | 
| 2146 | apply (subst density_max_0) | |
| 56996 | 2147 | apply (subst nn_integral_density) | 
| 2148 | apply (simp_all add: AE_count_space nn_integral_density) | |
| 2149 | apply (subst nn_integral_count_space ) | |
| 57418 | 2150 | apply (auto intro!: setsum.cong simp: max_def ereal_zero_less_0_iff) | 
| 47694 | 2151 | apply (rule finite_subset) | 
| 2152 | prefer 2 | |
| 2153 | apply assumption | |
| 2154 | apply auto | |
| 2155 | done | |
| 2156 | ||
| 56996 | 2157 | lemma nn_integral_point_measure_finite: | 
| 47694 | 2158 | "finite A \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> f a) \<Longrightarrow> (\<And>a. a \<in> A \<Longrightarrow> 0 \<le> g a) \<Longrightarrow> | 
| 56996 | 2159 | integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)" | 
| 57418 | 2160 | by (subst nn_integral_point_measure) (auto intro!: setsum.mono_neutral_left simp: less_le) | 
| 47694 | 2161 | |
| 56994 | 2162 | subsubsection {* Uniform measure *}
 | 
| 47694 | 2163 | |
| 2164 | definition "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)" | |
| 2165 | ||
| 2166 | lemma | |
| 2167 | shows sets_uniform_measure[simp]: "sets (uniform_measure M A) = sets M" | |
| 2168 | and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" | |
| 2169 | by (auto simp: uniform_measure_def) | |
| 2170 | ||
| 2171 | lemma emeasure_uniform_measure[simp]: | |
| 2172 | assumes A: "A \<in> sets M" and B: "B \<in> sets M" | |
| 2173 | shows "emeasure (uniform_measure M A) B = emeasure M (A \<inter> B) / emeasure M A" | |
| 2174 | proof - | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2175 | from A B have "emeasure (uniform_measure M A) B = (\<integral>\<^sup>+x. (1 / emeasure M A) * indicator (A \<inter> B) x \<partial>M)" | 
| 47694 | 2176 | by (auto simp add: uniform_measure_def emeasure_density split: split_indicator | 
| 56996 | 2177 | intro!: nn_integral_cong) | 
| 47694 | 2178 | also have "\<dots> = emeasure M (A \<inter> B) / emeasure M A" | 
| 2179 | using A B | |
| 56996 | 2180 | by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int emeasure_nonneg) | 
| 47694 | 2181 | finally show ?thesis . | 
| 2182 | qed | |
| 2183 | ||
| 2184 | lemma measure_uniform_measure[simp]: | |
| 2185 | assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>" and B: "B \<in> sets M" | |
| 2186 | shows "measure (uniform_measure M A) B = measure M (A \<inter> B) / measure M A" | |
| 2187 | using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A | |
| 2188 | by (cases "emeasure M A" "emeasure M (A \<inter> B)" rule: ereal2_cases) (simp_all add: measure_def) | |
| 2189 | ||
| 56994 | 2190 | subsubsection {* Uniform count measure *}
 | 
| 47694 | 2191 | |
| 2192 | definition "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)" | |
| 2193 | ||
| 2194 | lemma | |
| 2195 | shows space_uniform_count_measure: "space (uniform_count_measure A) = A" | |
| 2196 | and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" | |
| 2197 | unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) | |
| 2198 | ||
| 2199 | lemma emeasure_uniform_count_measure: | |
| 2200 | "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (uniform_count_measure A) X = card X / card A" | |
| 2201 | by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def) | |
| 2202 | ||
| 2203 | lemma measure_uniform_count_measure: | |
| 2204 | "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> measure (uniform_count_measure A) X = card X / card A" | |
| 2205 | by (simp add: real_eq_of_nat emeasure_point_measure_finite uniform_count_measure_def measure_def) | |
| 2206 | ||
| 35748 | 2207 | end |