| author | paulson <lp15@cam.ac.uk> | 
| Wed, 09 Dec 2015 17:35:22 +0000 | |
| changeset 61810 | 3c5040d5694a | 
| parent 61808 | fc1556774cfe | 
| child 61942 | f02b26f7d39d | 
| 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 | |
| 61808 | 6 | section \<open>Lebesgue Integration for Nonnegative Functions\<close> | 
| 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 | ||
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 12 | lemma infinite_countable_subset': | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 13 | assumes X: "infinite X" shows "\<exists>C\<subseteq>X. countable C \<and> infinite C" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 14 | proof - | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 15 | from infinite_countable_subset[OF X] guess f .. | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 16 | then show ?thesis | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 17 | by (intro exI[of _ "range f"]) (auto simp: range_inj_infinite) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 18 | qed | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 19 | |
| 56949 | 20 | lemma indicator_less_ereal[simp]: | 
| 21 | "indicator A x \<le> (indicator B x::ereal) \<longleftrightarrow> (x \<in> A \<longrightarrow> x \<in> B)" | |
| 22 | 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 | 23 | |
| 56994 | 24 | subsection "Simple function" | 
| 35582 | 25 | |
| 61808 | 26 | text \<open> | 
| 38656 | 27 | |
| 56996 | 28 | Our simple functions are not restricted to nonnegative real numbers. Instead | 
| 38656 | 29 | they are just functions with a finite range and are measurable when singleton | 
| 30 | sets are measurable. | |
| 35582 | 31 | |
| 61808 | 32 | \<close> | 
| 38656 | 33 | |
| 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 | 34 | definition "simple_function M g \<longleftrightarrow> | 
| 38656 | 35 | finite (g ` space M) \<and> | 
| 36 |     (\<forall>x \<in> g ` space M. g -` {x} \<inter> space M \<in> sets M)"
 | |
| 36624 | 37 | |
| 47694 | 38 | 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 | 39 | assumes "simple_function M g" | 
| 40875 | 40 | shows "finite (g ` space M)" and "g -` X \<inter> space M \<in> sets M" | 
| 40871 | 41 | proof - | 
| 42 | show "finite (g ` space M)" | |
| 43 | using assms unfolding simple_function_def by auto | |
| 40875 | 44 | have "g -` X \<inter> space M = g -` (X \<inter> g`space M) \<inter> space M" by auto | 
| 45 |   also have "\<dots> = (\<Union>x\<in>X \<inter> g`space M. g-`{x} \<inter> space M)" by auto
 | |
| 46 | 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 | 47 | by (auto simp del: UN_simps simp: simple_function_def) | 
| 40871 | 48 | qed | 
| 36624 | 49 | |
| 56949 | 50 | lemma measurable_simple_function[measurable_dest]: | 
| 51 | "simple_function M f \<Longrightarrow> f \<in> measurable M (count_space UNIV)" | |
| 52 | unfolding simple_function_def measurable_def | |
| 53 | proof safe | |
| 54 |   fix A assume "finite (f ` space M)" "\<forall>x\<in>f ` space M. f -` {x} \<inter> space M \<in> sets M"
 | |
| 55 |   then have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) \<in> sets M"
 | |
| 56 | by (intro sets.finite_UN) auto | |
| 57 |   also have "(\<Union>x\<in>f ` space M. if x \<in> A then f -` {x} \<inter> space M else {}) = f -` A \<inter> space M"
 | |
| 58 | by (auto split: split_if_asm) | |
| 59 | finally show "f -` A \<inter> space M \<in> sets M" . | |
| 60 | qed simp | |
| 61 | ||
| 62 | lemma borel_measurable_simple_function: | |
| 63 | "simple_function M f \<Longrightarrow> f \<in> borel_measurable M" | |
| 64 | by (auto dest!: measurable_simple_function simp: measurable_def) | |
| 65 | ||
| 47694 | 66 | lemma simple_function_measurable2[intro]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 67 | 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 | 68 | 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 | 69 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 70 | 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 | 71 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 72 | 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 | 73 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 74 | |
| 47694 | 75 | lemma simple_function_indicator_representation: | 
| 43920 | 76 | 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 | 77 | assumes f: "simple_function M f" and x: "x \<in> space M" | 
| 38656 | 78 |   shows "f x = (\<Sum>y \<in> f ` space M. y * indicator (f -` {y} \<inter> space M) x)"
 | 
| 79 | (is "?l = ?r") | |
| 80 | proof - | |
| 38705 | 81 | have "?r = (\<Sum>y \<in> f ` space M. | 
| 38656 | 82 |     (if y = f x then y * indicator (f -` {y} \<inter> space M) x else 0))"
 | 
| 57418 | 83 | by (auto intro!: setsum.cong) | 
| 38656 | 84 |   also have "... =  f x *  indicator (f -` {f x} \<inter> space M) x"
 | 
| 57418 | 85 | using assms by (auto dest: simple_functionD simp: setsum.delta) | 
| 38656 | 86 | also have "... = f x" using x by (auto simp: indicator_def) | 
| 87 | finally show ?thesis by auto | |
| 88 | qed | |
| 36624 | 89 | |
| 47694 | 90 | lemma simple_function_notspace: | 
| 43920 | 91 | "simple_function M (\<lambda>x. h x * indicator (- space M) x::ereal)" (is "simple_function M ?h") | 
| 35692 | 92 | proof - | 
| 38656 | 93 |   have "?h ` space M \<subseteq> {0}" unfolding indicator_def by auto
 | 
| 94 | hence [simp, intro]: "finite (?h ` space M)" by (auto intro: finite_subset) | |
| 95 |   have "?h -` {0} \<inter> space M = space M" by auto
 | |
| 96 | thus ?thesis unfolding simple_function_def by auto | |
| 97 | qed | |
| 98 | ||
| 47694 | 99 | lemma simple_function_cong: | 
| 38656 | 100 | 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 | 101 | shows "simple_function M f \<longleftrightarrow> simple_function M g" | 
| 38656 | 102 | proof - | 
| 103 | have "f ` space M = g ` space M" | |
| 104 |     "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
 | |
| 105 | using assms by (auto intro!: image_eqI) | |
| 106 | thus ?thesis unfolding simple_function_def using assms by simp | |
| 107 | qed | |
| 108 | ||
| 47694 | 109 | 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 | 110 | 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 | 111 | 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 | 112 | 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 | 113 | |
| 47694 | 114 | lemma simple_function_borel_measurable: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 115 |   fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
 | 
| 38656 | 116 | 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 | 117 | shows "simple_function M f" | 
| 38656 | 118 | using assms unfolding simple_function_def | 
| 119 | by (auto intro: borel_measurable_vimage) | |
| 120 | ||
| 56949 | 121 | lemma simple_function_eq_measurable: | 
| 43920 | 122 | fixes f :: "'a \<Rightarrow> ereal" | 
| 56949 | 123 | shows "simple_function M f \<longleftrightarrow> finite (f`space M) \<and> f \<in> measurable M (count_space UNIV)" | 
| 124 | 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 | 125 | by (fastforce simp: simple_function_def) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 126 | |
| 47694 | 127 | 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 | 128 | "simple_function M (\<lambda>x. c)" | 
| 38656 | 129 | by (auto intro: finite_subset simp: simple_function_def) | 
| 47694 | 130 | 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 | 131 | 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 | 132 | shows "simple_function M (g \<circ> f)" | 
| 38656 | 133 | unfolding simple_function_def | 
| 134 | proof safe | |
| 135 | show "finite ((g \<circ> f) ` space M)" | |
| 56154 
f0a927235162
more complete set of lemmas wrt. image and composition
 haftmann parents: 
54611diff
changeset | 136 | using assms unfolding simple_function_def by (auto simp: image_comp [symmetric]) | 
| 38656 | 137 | next | 
| 138 | fix x assume "x \<in> space M" | |
| 139 |   let ?G = "g -` {g (f x)} \<inter> (f`space M)"
 | |
| 140 |   have *: "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M =
 | |
| 141 |     (\<Union>x\<in>?G. f -` {x} \<inter> space M)" by auto
 | |
| 142 |   show "(g \<circ> f) -` {(g \<circ> f) x} \<inter> space M \<in> sets M"
 | |
| 143 | using assms unfolding simple_function_def * | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 144 | by (rule_tac sets.finite_UN) auto | 
| 38656 | 145 | qed | 
| 146 | ||
| 47694 | 147 | lemma simple_function_indicator[intro, simp]: | 
| 38656 | 148 | 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 | 149 | shows "simple_function M (indicator A)" | 
| 35692 | 150 | proof - | 
| 38656 | 151 |   have "indicator A ` space M \<subseteq> {0, 1}" (is "?S \<subseteq> _")
 | 
| 152 | by (auto simp: indicator_def) | |
| 153 | hence "finite ?S" by (rule finite_subset) simp | |
| 154 | moreover have "- A \<inter> space M = space M - A" by auto | |
| 155 | ultimately show ?thesis unfolding simple_function_def | |
| 46905 | 156 | using assms by (auto simp: indicator_def [abs_def]) | 
| 35692 | 157 | qed | 
| 158 | ||
| 47694 | 159 | 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 | 160 | 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 | 161 | 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 | 162 | shows "simple_function M (\<lambda>x. (f x, g x))" (is "simple_function M ?p") | 
| 38656 | 163 | unfolding simple_function_def | 
| 164 | proof safe | |
| 165 | show "finite (?p ` space M)" | |
| 166 | using assms unfolding simple_function_def | |
| 167 | by (rule_tac finite_subset[of _ "f`space M \<times> g`space M"]) auto | |
| 168 | next | |
| 169 | fix x assume "x \<in> space M" | |
| 170 |   have "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M =
 | |
| 171 |       (f -` {f x} \<inter> space M) \<inter> (g -` {g x} \<inter> space M)"
 | |
| 172 | by auto | |
| 61808 | 173 |   with \<open>x \<in> space M\<close> show "(\<lambda>x. (f x, g x)) -` {(f x, g x)} \<inter> space M \<in> sets M"
 | 
| 38656 | 174 | using assms unfolding simple_function_def by auto | 
| 175 | qed | |
| 35692 | 176 | |
| 47694 | 177 | 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 | 178 | 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 | 179 | shows "simple_function M (\<lambda>x. g (f x))" | 
| 38656 | 180 | using simple_function_compose[OF assms, of g] | 
| 181 | by (simp add: comp_def) | |
| 35582 | 182 | |
| 47694 | 183 | 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 | 184 | 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 | 185 | shows "simple_function M (\<lambda>x. h (f x) (g x))" | 
| 38656 | 186 | 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 | 187 | have "simple_function M ((\<lambda>(x, y). h x y) \<circ> (\<lambda>x. (f x, g x)))" | 
| 38656 | 188 | using assms by auto | 
| 189 | thus ?thesis by (simp_all add: comp_def) | |
| 190 | qed | |
| 35582 | 191 | |
| 47694 | 192 | lemmas simple_function_add[intro, simp] = simple_function_compose2[where h="op +"] | 
| 38656 | 193 | and simple_function_diff[intro, simp] = simple_function_compose2[where h="op -"] | 
| 194 | and simple_function_uminus[intro, simp] = simple_function_compose[where g="uminus"] | |
| 195 | and simple_function_mult[intro, simp] = simple_function_compose2[where h="op *"] | |
| 196 | and simple_function_div[intro, simp] = simple_function_compose2[where h="op /"] | |
| 197 | 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 | 198 | and simple_function_max[intro, simp] = simple_function_compose2[where h=max] | 
| 38656 | 199 | |
| 47694 | 200 | 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 | 201 | 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 | 202 | shows "simple_function M (\<lambda>x. \<Sum>i\<in>P. f i x)" | 
| 38656 | 203 | proof cases | 
| 204 | assume "finite P" from this assms show ?thesis by induct auto | |
| 205 | qed auto | |
| 35582 | 206 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 207 | lemma simple_function_ereal[intro, simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 208 | fixes f g :: "'a \<Rightarrow> real" assumes sf: "simple_function M f" | 
| 56949 | 209 | shows "simple_function M (\<lambda>x. ereal (f x))" | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 210 | by (rule simple_function_compose1[OF sf]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 211 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 212 | 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 | 213 | fixes f g :: "'a \<Rightarrow> nat" assumes sf: "simple_function M f" | 
| 56949 | 214 | shows "simple_function M (\<lambda>x. real (f x))" | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 215 | by (rule simple_function_compose1[OF sf]) | 
| 35582 | 216 | |
| 47694 | 217 | lemma borel_measurable_implies_simple_function_sequence: | 
| 43920 | 218 | fixes u :: "'a \<Rightarrow> ereal" | 
| 38656 | 219 | 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 | 220 | 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 | 221 | (\<forall>x. (SUP i. f i x) = max 0 (u x)) \<and> (\<forall>i x. 0 \<le> f i x)" | 
| 35582 | 222 | proof - | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 223 | def f \<equiv> "\<lambda>x i. if real i \<le> u x then i * 2 ^ i else nat(floor (real_of_ereal (u x) * 2 ^ i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 224 |   { 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 | 225 | proof (split split_if, intro conjI impI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 226 | assume "\<not> real j \<le> u x" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 227 | then have "nat(floor (real_of_ereal (u x) * 2 ^ j)) \<le> nat(floor (j * 2 ^ j))" | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 228 | by (cases "u x") (auto intro!: nat_mono floor_mono) | 
| 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 229 | moreover have "real (nat(floor (j * 2 ^ j))) \<le> j * 2^j" | 
| 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 230 | by linarith | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 231 | ultimately show "nat(floor (real_of_ereal (u x) * 2 ^ j)) \<le> j * 2 ^ j" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 232 | unfolding of_nat_le_iff by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 233 | qed auto } | 
| 38656 | 234 | note f_upper = this | 
| 35582 | 235 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 236 | have real_f: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 237 | "\<And>i x. real (f x i) = (if real i \<le> u x then i * 2 ^ i else real (nat(floor (real_of_ereal (u x) * 2 ^ i))))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 238 | unfolding f_def by auto | 
| 35582 | 239 | |
| 46731 | 240 | 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 | 241 | show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 242 | 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 | 243 | fix i | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 244 | 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 | 245 | proof (intro simple_function_borel_measurable) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 246 | 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 | 247 | 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 | 248 |       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 | 249 | using f_upper[of _ i] by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 250 | 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 | 251 | by (rule finite_subset) auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 252 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 253 | then show "simple_function M (?g i)" | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 254 | by (auto) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 255 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 256 | show "incseq ?g" | 
| 43920 | 257 | 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 | 258 | fix x and i :: nat | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 259 | 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 | 260 | proof ((split split_if)+, intro conjI impI) | 
| 43920 | 261 | assume "ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 262 | then show "i * 2 ^ i * 2 \<le> nat(floor (real_of_ereal (u x) * 2 ^ Suc i))" | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 263 | by (cases "u x") (auto intro!: le_nat_floor) | 
| 38656 | 264 | next | 
| 43920 | 265 | assume "\<not> ereal (real i) \<le> u x" "ereal (real (Suc i)) \<le> u x" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 266 | then show "nat(floor (real_of_ereal (u x) * 2 ^ i)) * 2 \<le> Suc i * 2 ^ Suc i" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 267 | by (cases "u x") auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 268 | next | 
| 43920 | 269 | assume "\<not> ereal (real i) \<le> u x" "\<not> ereal (real (Suc i)) \<le> u x" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 270 | have "nat(floor (real_of_ereal (u x) * 2 ^ i)) * 2 = nat(floor (real_of_ereal (u x) * 2 ^ i)) * nat(floor(2::real))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 271 | by simp | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 272 | also have "\<dots> \<le> nat(floor (real_of_ereal (u x) * 2 ^ i * 2))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 273 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 274 | assume "0 \<le> u x" then show ?thesis | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 275 | by (intro le_mult_nat_floor) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 276 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 277 | assume "\<not> 0 \<le> u x" then show ?thesis | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 278 | by (cases "u x") (auto simp: nat_floor_neg mult_nonpos_nonneg) | 
| 38656 | 279 | qed | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 280 | also have "\<dots> = nat(floor (real_of_ereal (u x) * 2 ^ Suc i))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 281 | by (simp add: ac_simps) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 282 | finally show "nat(floor (real_of_ereal (u x) * 2 ^ i)) * 2 \<le> nat(floor (real_of_ereal (u x) * 2 ^ Suc i))" . | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 283 | qed simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 284 | 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 | 285 | by (auto simp: field_simps) | 
| 35582 | 286 | qed | 
| 38656 | 287 | next | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 288 | fix x show "(SUP i. ?g i x) = max 0 (u x)" | 
| 51000 | 289 | proof (rule SUP_eqI) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 290 | fix i show "?g i x \<le> max 0 (u x)" unfolding max_def f_def | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 291 | by (cases "u x") (auto simp: field_simps nat_floor_neg mult_nonpos_nonneg) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 292 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 293 | 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 | 294 | 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 | 295 | 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 | 296 | show "max 0 (u x) \<le> y" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 297 | proof (cases y) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 298 | case (real r) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 299 | with * have *: "\<And>i. f x i \<le> r * 2^i" by (auto simp: divide_le_eq) | 
| 44666 | 300 | from reals_Archimedean2[of r] * have "u x \<noteq> \<infinity>" by (auto simp: f_def) (metis less_le_not_le) | 
| 43920 | 301 | 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 | 302 | then guess p .. note ux = this | 
| 44666 | 303 | 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 | 304 | have "p \<le> r" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 305 | proof (rule ccontr) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 306 | assume "\<not> p \<le> r" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 307 | with LIMSEQ_inverse_realpow_zero[unfolded LIMSEQ_iff, rule_format, of 2 "p - r"] | 
| 56536 | 308 | 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 | 309 | 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 | 310 | moreover | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 311 | have "real (nat(floor (p * 2 ^ max N m))) \<le> r * 2 ^ max N m" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 312 | using *[of "max N m"] m unfolding real_f using ux | 
| 56536 | 313 | 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 | 314 | then have "p * 2 ^ max N m - 1 < r * 2 ^ max N m" | 
| 59587 
8ea7b22525cb
Removed the obsolete functions "natfloor" and "natceiling"
 nipkow parents: 
59452diff
changeset | 315 | by linarith | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 316 | ultimately show False by auto | 
| 38656 | 317 | qed | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 318 | then show "max 0 (u x) \<le> y" using real ux by simp | 
| 61808 | 319 | qed (insert \<open>0 \<le> y\<close>, auto) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 320 | qed | 
| 56571 
f4635657d66f
added divide_nonneg_nonneg and co; made it a simp rule
 hoelzl parents: 
56537diff
changeset | 321 | qed auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 322 | qed | 
| 35582 | 323 | |
| 47694 | 324 | lemma borel_measurable_implies_simple_function_sequence': | 
| 43920 | 325 | fixes u :: "'a \<Rightarrow> ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 326 | assumes u: "u \<in> borel_measurable M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 327 | 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 | 328 | "\<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 | 329 | 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 | 330 | |
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 331 | 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 | 332 | fixes u :: "'a \<Rightarrow> ereal" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 333 | assumes u: "simple_function M u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 334 | 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 | 335 | 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 | 336 | 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 | 337 | 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 | 338 | shows "P u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 339 | proof (rule cong) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 340 |   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 | 341 | proof eventually_elim | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 342 | fix x assume x: "x \<in> space M" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 343 | from simple_function_indicator_representation[OF u x] | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 344 |     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 | 345 | qed | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 346 | next | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 347 | from u have "finite (u ` space M)" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 348 | unfolding simple_function_def by auto | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 349 |   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 | 350 | proof induct | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 351 | case empty show ?case | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 352 |       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 | 353 | 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 | 354 | next | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 355 |   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 | 356 | apply (subst simple_function_cong) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 357 | apply (rule simple_function_indicator_representation[symmetric]) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 358 | apply (auto intro: u) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 359 | done | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 360 | qed fact | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 361 | |
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 362 | 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 | 363 | fixes u :: "'a \<Rightarrow> ereal" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 364 | 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 | 365 | 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 | 366 | 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 | 367 | 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 | 368 | 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 | 369 | shows "P u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 370 | proof - | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 371 | show ?thesis | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 372 | proof (rule cong) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 373 | fix x assume x: "x \<in> space M" | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 374 | from simple_function_indicator_representation[OF u x] | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 375 |     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 | 376 | next | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 377 |     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 | 378 | apply (subst simple_function_cong) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 379 | apply (rule simple_function_indicator_representation[symmetric]) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 380 | apply (auto intro: u) | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 381 | done | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 382 | next | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 383 | |
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 384 | 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 | 385 | unfolding simple_function_def by auto | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 386 |     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 | 387 | proof induct | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 388 | case empty show ?case | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 389 |         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 | 390 | next | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 391 | 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 | 392 |       { 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 | 393 |           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 | 394 | 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 | 395 | 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 | 396 | 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 | 397 | 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 | 398 | qed | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 399 | qed fact | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 400 | qed | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 401 | |
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 402 | 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 | 403 | fixes u :: "'a \<Rightarrow> ereal" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 404 | 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 | 405 | 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 | 406 | 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 | 407 | 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 | 408 | 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 | 409 | 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 | 410 | shows "P u" | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 411 | using u | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 412 | 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 | 413 | 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 | 414 | 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 | 415 | 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 | 416 | 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 | 417 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 418 | 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 | 419 | using U by (auto simp: image_iff eq_commute) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 420 | |
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 421 | 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 | 422 | by (simp add: borel_measurable_simple_function) | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 423 | |
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 424 | show "P u" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 425 | unfolding u_eq | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 426 | proof (rule seq) | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 427 | fix i show "P (U i)" | 
| 61808 | 428 | using \<open>simple_function M (U i)\<close> nn[of i] not_inf[of _ i] | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 429 | 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 | 430 | 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 | 431 | show ?case | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 432 | proof cases | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 433 |         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 | 434 | 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 | 435 |           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 | 436 | (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 | 437 | next | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 438 |         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 | 439 | 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 | 440 | 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 | 441 | by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 442 | 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 | 443 | by auto | 
| 61808 | 444 | from x mult(5)[OF \<open>x \<in> space M\<close>] mult(1) mult(3)[of x] have "c < \<infinity>" | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 445 | by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 446 | 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 | 447 | show ?thesis | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 448 | 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 | 449 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 450 | 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 | 451 | qed fact+ | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 452 | qed | 
| 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 453 | |
| 47694 | 454 | lemma simple_function_If_set: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 455 | 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 | 456 | 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 | 457 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 458 |   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 | 459 | show ?thesis unfolding simple_function_def | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 460 | proof safe | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 461 | 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 | 462 | from finite_subset[OF this] assms | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 463 | 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 | 464 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 465 | fix x assume "x \<in> space M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 466 |     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 | 467 | 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 | 468 | 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 | 469 | 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 | 470 | 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 | 471 | 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 | 472 |     show "?IF -` {?IF x} \<inter> space M \<in> sets M" unfolding * using A by auto
 | 
| 35582 | 473 | qed | 
| 474 | qed | |
| 475 | ||
| 47694 | 476 | lemma simple_function_If: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 477 |   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 | 478 | shows "simple_function M (\<lambda>x. if P x then f x else g x)" | 
| 35582 | 479 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 480 |   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 | 481 |   with simple_function_If_set[OF sf, of "{x. P x}"] P show ?thesis by simp
 | 
| 38656 | 482 | qed | 
| 483 | ||
| 47694 | 484 | 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 | 485 | 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 | 486 | 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 | 487 | 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 | 488 | using assms unfolding simple_function_def by auto | 
| 39092 | 489 | |
| 47694 | 490 | lemma simple_function_comp: | 
| 491 | 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 | 492 | 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 | 493 | shows "simple_function M (\<lambda>x. f (T x))" | 
| 41661 | 494 | proof (intro simple_function_def[THEN iffD2] conjI ballI) | 
| 495 | have "(\<lambda>x. f (T x)) ` space M \<subseteq> f ` space M'" | |
| 496 | using T unfolding measurable_def by auto | |
| 497 | 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 | 498 | using f unfolding simple_function_def by (auto intro: finite_subset) | 
| 41661 | 499 | fix i assume i: "i \<in> (\<lambda>x. f (T x)) ` space M" | 
| 500 | then have "i \<in> f ` space M'" | |
| 501 | using T unfolding measurable_def by auto | |
| 502 |   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 | 503 | using f unfolding simple_function_def by auto | 
| 41661 | 504 |   then have "T -` (f -` {i} \<inter> space M') \<inter> space M \<in> sets M"
 | 
| 505 | using T unfolding measurable_def by auto | |
| 506 |   also have "T -` (f -` {i} \<inter> space M') \<inter> space M = (\<lambda>x. f (T x)) -` {i} \<inter> space M"
 | |
| 507 | using T unfolding measurable_def by auto | |
| 508 |   finally show "(\<lambda>x. f (T x)) -` {i} \<inter> space M \<in> sets M" .
 | |
| 40859 | 509 | qed | 
| 510 | ||
| 56994 | 511 | subsection "Simple integral" | 
| 38656 | 512 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 513 | 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 | 514 |   "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 | 515 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 516 | syntax | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 517 |   "_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 | 518 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 519 | translations | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 520 | "\<integral>\<^sup>S x. f \<partial>M" == "CONST simple_integral M (%x. f)" | 
| 35582 | 521 | |
| 47694 | 522 | lemma simple_integral_cong: | 
| 38656 | 523 | 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 | 524 | shows "integral\<^sup>S M f = integral\<^sup>S M g" | 
| 38656 | 525 | proof - | 
| 526 | have "f ` space M = g ` space M" | |
| 527 |     "\<And>x. f -` {x} \<inter> space M = g -` {x} \<inter> space M"
 | |
| 528 | using assms by (auto intro!: image_eqI) | |
| 529 | thus ?thesis unfolding simple_integral_def by simp | |
| 530 | qed | |
| 531 | ||
| 47694 | 532 | 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 | 533 | "(\<integral>\<^sup>Sx. c \<partial>M) = c * (emeasure M) (space M)" | 
| 38656 | 534 | proof (cases "space M = {}")
 | 
| 535 | case True thus ?thesis unfolding simple_integral_def by simp | |
| 536 | next | |
| 537 |   case False hence "(\<lambda>x. c) ` space M = {c}" by auto
 | |
| 538 | thus ?thesis unfolding simple_integral_def by simp | |
| 35582 | 539 | qed | 
| 540 | ||
| 47694 | 541 | lemma simple_function_partition: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 542 | assumes f: "simple_function M f" and g: "simple_function M g" | 
| 56949 | 543 | assumes sub: "\<And>x y. x \<in> space M \<Longrightarrow> y \<in> space M \<Longrightarrow> g x = g y \<Longrightarrow> f x = f y" | 
| 544 | assumes v: "\<And>x. x \<in> space M \<Longrightarrow> f x = v (g x)" | |
| 545 |   shows "integral\<^sup>S M f = (\<Sum>y\<in>g ` space M. v y * emeasure M {x\<in>space M. g x = y})"
 | |
| 546 | (is "_ = ?r") | |
| 547 | proof - | |
| 548 | from f g have [simp]: "finite (f`space M)" "finite (g`space M)" | |
| 549 | by (auto simp: simple_function_def) | |
| 550 | from f g have [measurable]: "f \<in> measurable M (count_space UNIV)" "g \<in> measurable M (count_space UNIV)" | |
| 551 | by (auto intro: measurable_simple_function) | |
| 35582 | 552 | |
| 56949 | 553 |   { fix y assume "y \<in> space M"
 | 
| 554 |     then have "f ` space M \<inter> {i. \<exists>x\<in>space M. i = f x \<and> g y = g x} = {v (g y)}"
 | |
| 555 | by (auto cong: sub simp: v[symmetric]) } | |
| 556 | note eq = this | |
| 35582 | 557 | |
| 56949 | 558 | have "integral\<^sup>S M f = | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 559 | (\<Sum>y\<in>f`space M. y * (\<Sum>z\<in>g`space M. | 
| 56949 | 560 |       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))"
 | 
| 561 | unfolding simple_integral_def | |
| 59002 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 hoelzl parents: 
59000diff
changeset | 562 | proof (safe intro!: setsum.cong ereal_right_mult_cong) | 
| 56949 | 563 | fix y assume y: "y \<in> space M" "f y \<noteq> 0" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 564 |     have [simp]: "g ` space M \<inter> {z. \<exists>x\<in>space M. f y = f x \<and> z = g x} =
 | 
| 56949 | 565 |         {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
 | 
| 566 | by auto | |
| 567 |     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}) =
 | |
| 568 |         f -` {f y} \<inter> space M"
 | |
| 569 | by (auto simp: eq_commute cong: sub rev_conj_cong) | |
| 570 | have "finite (g`space M)" by simp | |
| 571 |     then have "finite {z. \<exists>x\<in>space M. f y = f x \<and> z = g x}"
 | |
| 572 | by (rule rev_finite_subset) auto | |
| 573 |     then show "emeasure M (f -` {f y} \<inter> space M) =
 | |
| 574 |       (\<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 | 575 | apply (simp add: setsum.If_cases) | 
| 56949 | 576 | apply (subst setsum_emeasure) | 
| 577 | apply (auto simp: disjoint_family_on_def eq) | |
| 578 | done | |
| 38656 | 579 | qed | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 580 | also have "\<dots> = (\<Sum>y\<in>f`space M. (\<Sum>z\<in>g`space M. | 
| 56949 | 581 |       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 | 582 | by (auto intro!: setsum.cong simp: setsum_ereal_right_distrib emeasure_nonneg) | 
| 56949 | 583 | also have "\<dots> = ?r" | 
| 57418 | 584 | by (subst setsum.commute) | 
| 585 | (auto intro!: setsum.cong simp: setsum.If_cases scaleR_setsum_right[symmetric] eq) | |
| 56949 | 586 | finally show "integral\<^sup>S M f = ?r" . | 
| 35582 | 587 | qed | 
| 588 | ||
| 47694 | 589 | lemma simple_integral_add[simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 590 | 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 | 591 | shows "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = integral\<^sup>S M f + integral\<^sup>S M g" | 
| 35582 | 592 | proof - | 
| 56949 | 593 | have "(\<integral>\<^sup>Sx. f x + g x \<partial>M) = | 
| 594 |     (\<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})"
 | |
| 595 | by (intro simple_function_partition) (auto intro: f g) | |
| 596 |   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}) +
 | |
| 597 |     (\<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 | 598 | using assms(2,4) by (auto intro!: setsum.cong ereal_left_distrib simp: setsum.distrib[symmetric]) | 
| 56949 | 599 |   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)"
 | 
| 600 | by (intro simple_function_partition[symmetric]) (auto intro: f g) | |
| 601 |   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)"
 | |
| 602 | by (intro simple_function_partition[symmetric]) (auto intro: f g) | |
| 603 | finally show ?thesis . | |
| 35582 | 604 | qed | 
| 605 | ||
| 47694 | 606 | lemma simple_integral_setsum[simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 607 | 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 | 608 | 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 | 609 | shows "(\<integral>\<^sup>Sx. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>S M (f i))" | 
| 38656 | 610 | proof cases | 
| 611 | assume "finite P" | |
| 612 | from this assms show ?thesis | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 613 | by induct (auto simp: simple_function_setsum simple_integral_add setsum_nonneg) | 
| 38656 | 614 | qed auto | 
| 615 | ||
| 47694 | 616 | lemma simple_integral_mult[simp]: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 617 | 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 | 618 | shows "(\<integral>\<^sup>Sx. c * f x \<partial>M) = c * integral\<^sup>S M f" | 
| 38656 | 619 | proof - | 
| 56949 | 620 |   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})"
 | 
| 621 | using f by (intro simple_function_partition) auto | |
| 622 | also have "\<dots> = c * integral\<^sup>S M f" | |
| 623 | using f unfolding simple_integral_def | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
57447diff
changeset | 624 | by (subst setsum_ereal_right_distrib) (auto simp: emeasure_nonneg mult.assoc Int_def conj_commute) | 
| 56949 | 625 | finally show ?thesis . | 
| 40871 | 626 | qed | 
| 627 | ||
| 47694 | 628 | lemma simple_integral_mono_AE: | 
| 56949 | 629 | assumes f[measurable]: "simple_function M f" and g[measurable]: "simple_function M g" | 
| 47694 | 630 | 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 | 631 | shows "integral\<^sup>S M f \<le> integral\<^sup>S M g" | 
| 40859 | 632 | proof - | 
| 56949 | 633 |   let ?\<mu> = "\<lambda>P. emeasure M {x\<in>space M. P x}"
 | 
| 634 | 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))" | |
| 635 | using f g by (intro simple_function_partition) auto | |
| 636 | 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))" | |
| 637 | proof (clarsimp intro!: setsum_mono) | |
| 40859 | 638 | fix x assume "x \<in> space M" | 
| 56949 | 639 | let ?M = "?\<mu> (\<lambda>y. f y = f x \<and> g y = g x)" | 
| 640 | show "f x * ?M \<le> g x * ?M" | |
| 641 | proof cases | |
| 642 | assume "?M \<noteq> 0" | |
| 643 | then have "0 < ?M" | |
| 644 | by (simp add: less_le emeasure_nonneg) | |
| 645 | also have "\<dots> \<le> ?\<mu> (\<lambda>y. f x \<le> g x)" | |
| 646 | using mono by (intro emeasure_mono_AE) auto | |
| 647 | finally have "\<not> \<not> f x \<le> g x" | |
| 648 | by (intro notI) auto | |
| 649 | then show ?thesis | |
| 650 | by (intro ereal_mult_right_mono) auto | |
| 651 | qed simp | |
| 40859 | 652 | qed | 
| 56949 | 653 | also have "\<dots> = integral\<^sup>S M g" | 
| 654 | using f g by (intro simple_function_partition[symmetric]) auto | |
| 655 | finally show ?thesis . | |
| 40859 | 656 | qed | 
| 657 | ||
| 47694 | 658 | 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 | 659 | assumes "simple_function M f" and "simple_function M g" | 
| 38656 | 660 | 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 | 661 | shows "integral\<^sup>S M f \<le> integral\<^sup>S M g" | 
| 41705 | 662 | using assms by (intro simple_integral_mono_AE) auto | 
| 35582 | 663 | |
| 47694 | 664 | lemma simple_integral_cong_AE: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 665 | assumes "simple_function M f" and "simple_function M g" | 
| 47694 | 666 | 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 | 667 | shows "integral\<^sup>S M f = integral\<^sup>S M g" | 
| 40859 | 668 | using assms by (auto simp: eq_iff intro!: simple_integral_mono_AE) | 
| 669 | ||
| 47694 | 670 | 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 | 671 | assumes sf: "simple_function M f" "simple_function M g" | 
| 47694 | 672 |   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 | 673 | shows "integral\<^sup>S M f = integral\<^sup>S M g" | 
| 40859 | 674 | proof (intro simple_integral_cong_AE sf AE_I) | 
| 47694 | 675 |   show "(emeasure M) {x\<in>space M. f x \<noteq> g x} = 0" by fact
 | 
| 40859 | 676 |   show "{x \<in> space M. f x \<noteq> g x} \<in> sets M"
 | 
| 677 | using sf[THEN borel_measurable_simple_function] by auto | |
| 678 | qed simp | |
| 679 | ||
| 47694 | 680 | lemma simple_integral_indicator: | 
| 56949 | 681 | assumes A: "A \<in> sets M" | 
| 49796 
182fa22e7ee8
introduce induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49795diff
changeset | 682 | 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 | 683 | shows "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) = | 
| 56949 | 684 |     (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
 | 
| 685 | proof - | |
| 686 |   have eq: "(\<lambda>x. (f x, indicator A x)) ` space M \<inter> {x. snd x = 1} = (\<lambda>x. (f x, 1::ereal))`A"
 | |
| 687 | using A[THEN sets.sets_into_space] by (auto simp: indicator_def image_iff split: split_if_asm) | |
| 688 |   have eq2: "\<And>x. f x \<notin> f ` A \<Longrightarrow> f -` {f x} \<inter> space M \<inter> A = {}"
 | |
| 689 | by (auto simp: image_iff) | |
| 690 | ||
| 691 | have "(\<integral>\<^sup>Sx. f x * indicator A x \<partial>M) = | |
| 692 |     (\<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})"
 | |
| 693 | using assms by (intro simple_function_partition) auto | |
| 694 | also have "\<dots> = (\<Sum>y\<in>(\<lambda>x. (f x, indicator A x::ereal))`space M. | |
| 695 |     if snd y = 1 then fst y * emeasure M (f -` {fst y} \<inter> space M \<inter> A) else 0)"
 | |
| 57418 | 696 | by (auto simp: indicator_def split: split_if_asm intro!: arg_cong2[where f="op *"] arg_cong2[where f=emeasure] setsum.cong) | 
| 56949 | 697 |   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 | 698 | using assms by (subst setsum.If_cases) (auto intro!: simple_functionD(1) simp: eq) | 
| 56949 | 699 |   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 | 700 | by (subst setsum.reindex [of fst]) (auto simp: inj_on_def) | 
| 56949 | 701 |   also have "\<dots> = (\<Sum>x \<in> f ` space M. x * emeasure M (f -` {x} \<inter> space M \<inter> A))"
 | 
| 702 | using A[THEN sets.sets_into_space] | |
| 57418 | 703 | by (intro setsum.mono_neutral_cong_left simple_functionD f) (auto simp: image_comp comp_def eq2) | 
| 56949 | 704 | finally show ?thesis . | 
| 38656 | 705 | qed | 
| 35582 | 706 | |
| 47694 | 707 | lemma simple_integral_indicator_only[simp]: | 
| 38656 | 708 | 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 | 709 | shows "integral\<^sup>S M (indicator A) = emeasure M A" | 
| 56949 | 710 | using simple_integral_indicator[OF assms, of "\<lambda>x. 1"] sets.sets_into_space[OF assms] | 
| 711 | by (simp_all add: image_constant_conv Int_absorb1 split: split_if_asm) | |
| 35582 | 712 | |
| 47694 | 713 | lemma simple_integral_null_set: | 
| 714 | 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 | 715 | shows "(\<integral>\<^sup>Sx. u x * indicator N x \<partial>M) = 0" | 
| 38656 | 716 | proof - | 
| 47694 | 717 | have "AE x in M. indicator N x = (0 :: ereal)" | 
| 61808 | 718 | using \<open>N \<in> null_sets M\<close> 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 | 719 | 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 | 720 | using assms apply (intro simple_integral_cong_AE) by auto | 
| 40859 | 721 | then show ?thesis by simp | 
| 38656 | 722 | qed | 
| 35582 | 723 | |
| 47694 | 724 | lemma simple_integral_cong_AE_mult_indicator: | 
| 725 | 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 | 726 | shows "integral\<^sup>S M f = (\<integral>\<^sup>Sx. f x * indicator S x \<partial>M)" | 
| 41705 | 727 | using assms by (intro simple_integral_cong_AE) auto | 
| 35582 | 728 | |
| 47694 | 729 | lemma simple_integral_cmult_indicator: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 730 | assumes A: "A \<in> sets M" | 
| 56949 | 731 | 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 | 732 | 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 | 733 | 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 | 734 | |
| 56996 | 735 | lemma simple_integral_nonneg: | 
| 47694 | 736 | 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 | 737 | 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 | 738 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 739 | 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 | 740 | 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 | 741 | then show ?thesis by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 742 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 743 | |
| 61808 | 744 | subsection \<open>Integral on nonnegative functions\<close> | 
| 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 | 745 | |
| 56996 | 746 | definition nn_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> ereal" ("integral\<^sup>N") where
 | 
| 747 |   "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 | 748 | |
| 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 | 749 | syntax | 
| 59357 
f366643536cd
allow line breaks in integral notation
 Andreas Lochbihler parents: 
59048diff
changeset | 750 |   "_nn_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> 'a measure \<Rightarrow> ereal" ("\<integral>\<^sup>+((2 _./ _)/ \<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 | 751 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 752 | translations | 
| 56996 | 753 | "\<integral>\<^sup>+x. f \<partial>M" == "CONST nn_integral M (\<lambda>x. f)" | 
| 40872 | 754 | |
| 57025 | 755 | lemma nn_integral_nonneg: "0 \<le> integral\<^sup>N M f" | 
| 56996 | 756 | by (auto intro!: SUP_upper2[of "\<lambda>x. 0"] simp: nn_integral_def le_fun_def) | 
| 40873 | 757 | |
| 58606 | 758 | lemma nn_integral_le_0[simp]: "integral\<^sup>N M f \<le> 0 \<longleftrightarrow> integral\<^sup>N M f = 0" | 
| 759 | using nn_integral_nonneg[of M f] by auto | |
| 760 | ||
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 761 | lemma nn_integral_not_less_0 [simp]: "\<not> nn_integral M f < 0" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 762 | by(simp add: not_less nn_integral_nonneg) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 763 | |
| 56996 | 764 | lemma nn_integral_not_MInfty[simp]: "integral\<^sup>N M f \<noteq> -\<infinity>" | 
| 765 | using nn_integral_nonneg[of M f] by auto | |
| 47694 | 766 | |
| 56996 | 767 | lemma nn_integral_def_finite: | 
| 768 |   "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 | 769 | (is "_ = SUPREMUM ?A ?f") | 
| 56996 | 770 | unfolding nn_integral_def | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 771 | proof (safe intro!: antisym SUP_least) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 772 | 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 | 773 |   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 | 774 | note gM = g(1)[THEN borel_measurable_simple_function] | 
| 50252 | 775 | have \<mu>_G_pos: "0 \<le> (emeasure M) ?G" using gM by auto | 
| 46731 | 776 | 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 | 777 | 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 | 778 | 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 | 779 | 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 | 780 | 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 | 781 | 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 | 782 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 783 | have g0: "?g 0 \<in> ?A" by (intro g_in_A) auto | 
| 47694 | 784 | assume "(emeasure M) ?G = 0" | 
| 785 | with gM have "AE x in M. x \<notin> ?G" | |
| 786 | 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 | 787 | with gM g show ?thesis | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 788 | 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 | 789 | (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 | 790 | next | 
| 50252 | 791 | 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 | 792 | 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 | 793 | proof (intro SUP_PInfty) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 794 | fix n :: nat | 
| 47694 | 795 | let ?y = "ereal (real n) / (if (emeasure M) ?G = \<infinity> then 1 else (emeasure M) ?G)" | 
| 50252 | 796 | 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 | 797 | then have "?g ?y \<in> ?A" by (rule g_in_A) | 
| 47694 | 798 | have "real n \<le> ?y * (emeasure M) ?G" | 
| 50252 | 799 | 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 | 800 | also have "\<dots> = (\<integral>\<^sup>Sx. ?y * indicator ?G x \<partial>M)" | 
| 61808 | 801 | using \<open>0 \<le> ?y\<close> \<open>?g ?y \<in> ?A\<close> gM | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 802 | by (subst simple_integral_cmult_indicator) auto | 
| 61808 | 803 | also have "\<dots> \<le> integral\<^sup>S M (?g ?y)" using \<open>?g ?y \<in> ?A\<close> gM | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 804 | 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 | 805 | finally show "\<exists>i\<in>?A. real n \<le> integral\<^sup>S M i" | 
| 61808 | 806 | using \<open>?g ?y \<in> ?A\<close> by blast | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 807 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 808 | then show ?thesis by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 809 | qed | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 810 | qed (auto intro: SUP_upper) | 
| 40873 | 811 | |
| 56996 | 812 | lemma nn_integral_mono_AE: | 
| 813 | assumes ae: "AE x in M. u x \<le> v x" shows "integral\<^sup>N M u \<le> integral\<^sup>N M v" | |
| 814 | unfolding nn_integral_def | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 815 | proof (safe intro!: SUP_mono) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 816 | 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 | 817 | from ae[THEN AE_E] guess N . note N = this | 
| 47694 | 818 | then have ae_N: "AE x in M. x \<notin> N" by (auto intro: AE_not_in) | 
| 46731 | 819 | let ?n = "\<lambda>x. n x * indicator (space M - N) x" | 
| 47694 | 820 | 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 | 821 | using n N ae_N by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 822 | moreover | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 823 |   { 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 | 824 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 825 | assume x: "x \<in> space M - N" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 826 | 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 | 827 | 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 | 828 | 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 | 829 | qed simp } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 830 | 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 | 831 | 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 | 832 | 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 | 833 |   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 | 834 | by force | 
| 38656 | 835 | qed | 
| 836 | ||
| 56996 | 837 | lemma nn_integral_mono: | 
| 838 | "(\<And>x. x \<in> space M \<Longrightarrow> u x \<le> v x) \<Longrightarrow> integral\<^sup>N M u \<le> integral\<^sup>N M v" | |
| 839 | by (auto intro: nn_integral_mono_AE) | |
| 40859 | 840 | |
| 60175 | 841 | lemma mono_nn_integral: "mono F \<Longrightarrow> mono (\<lambda>x. integral\<^sup>N M (F x))" | 
| 842 | by (auto simp add: mono_def le_fun_def intro!: nn_integral_mono) | |
| 843 | ||
| 56996 | 844 | lemma nn_integral_cong_AE: | 
| 845 | "AE x in M. u x = v x \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v" | |
| 846 | by (auto simp: eq_iff intro!: nn_integral_mono_AE) | |
| 40859 | 847 | |
| 56996 | 848 | lemma nn_integral_cong: | 
| 849 | "(\<And>x. x \<in> space M \<Longrightarrow> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v" | |
| 850 | by (auto intro: nn_integral_cong_AE) | |
| 40859 | 851 | |
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 852 | lemma nn_integral_cong_simp: | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 853 | "(\<And>x. x \<in> space M =simp=> u x = v x) \<Longrightarrow> integral\<^sup>N M u = integral\<^sup>N M v" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 854 | by (auto intro: nn_integral_cong simp: simp_implies_def) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 855 | |
| 56996 | 856 | lemma nn_integral_cong_strong: | 
| 857 | "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" | |
| 858 | 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 | 859 | |
| 56996 | 860 | lemma nn_integral_eq_simple_integral: | 
| 861 | 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 | 862 | proof - | 
| 46731 | 863 | 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 | 864 | 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 | 865 | 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 | 866 | by (auto simp: fun_eq_iff max_def split: split_indicator) | 
| 56996 | 867 | have "integral\<^sup>N M ?f \<le> integral\<^sup>S M ?f" using f' | 
| 868 | by (force intro!: SUP_least simple_integral_mono simp: le_fun_def nn_integral_def) | |
| 869 | moreover have "integral\<^sup>S M ?f \<le> integral\<^sup>N M ?f" | |
| 870 | unfolding nn_integral_def | |
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 871 | 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 | 872 | ultimately show ?thesis | 
| 56996 | 873 | 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 | 874 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 875 | |
| 56996 | 876 | lemma nn_integral_eq_simple_integral_AE: | 
| 877 | 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 | 878 | proof - | 
| 47694 | 879 | have "AE x in M. f x = max 0 (f x)" using f by (auto split: split_max) | 
| 56996 | 880 | with f have "integral\<^sup>N M f = integral\<^sup>S M (\<lambda>x. max 0 (f x))" | 
| 881 | by (simp cong: nn_integral_cong_AE simple_integral_cong_AE | |
| 882 | add: nn_integral_eq_simple_integral) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 883 | with assms show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 884 | 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 | 885 | qed | 
| 40873 | 886 | |
| 56996 | 887 | lemma nn_integral_SUP_approx: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 888 | 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 | 889 |   and u: "simple_function M u" "u \<le> (SUP i. f i)" "u`space M \<subseteq> {0..<\<infinity>}"
 | 
| 56996 | 890 | shows "integral\<^sup>S M u \<le> (SUP i. integral\<^sup>N M (f i))" (is "_ \<le> ?S") | 
| 43920 | 891 | proof (rule ereal_le_mult_one_interval) | 
| 56996 | 892 | have "0 \<le> (SUP i. integral\<^sup>N M (f i))" | 
| 893 | using f(3) by (auto intro!: SUP_upper2 nn_integral_nonneg) | |
| 894 | 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 | 895 | 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 | 896 | using u(3) by auto | 
| 43920 | 897 | fix a :: ereal assume "0 < a" "a < 1" | 
| 38656 | 898 | hence "a \<noteq> 0" by auto | 
| 46731 | 899 |   let ?B = "\<lambda>i. {x \<in> space M. a * u x \<le> f i x}"
 | 
| 38656 | 900 | have B: "\<And>i. ?B i \<in> sets M" | 
| 61808 | 901 | using f \<open>simple_function M u\<close>[THEN borel_measurable_simple_function] by auto | 
| 38656 | 902 | |
| 46731 | 903 | let ?uB = "\<lambda>i x. u x * indicator (?B i) x" | 
| 38656 | 904 | |
| 905 |   { fix i have "?B i \<subseteq> ?B (Suc i)"
 | |
| 906 | proof safe | |
| 907 | fix i x assume "a * u x \<le> f i x" | |
| 908 | also have "\<dots> \<le> f (Suc i) x" | |
| 61808 | 909 | using \<open>incseq f\<close>[THEN incseq_SucD] unfolding le_fun_def by auto | 
| 38656 | 910 | finally show "a * u x \<le> f (Suc i) x" . | 
| 911 | qed } | |
| 912 | note B_mono = this | |
| 35582 | 913 | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 914 | note B_u = sets.Int[OF u(1)[THEN simple_functionD(2)] B] | 
| 38656 | 915 | |
| 46731 | 916 |   let ?B' = "\<lambda>i n. (u -` {i} \<inter> space M) \<inter> ?B n"
 | 
| 47694 | 917 |   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 | 918 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 919 | fix i | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 920 | 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 | 921 | 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 | 922 |     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 | 923 | proof safe | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 924 | 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 | 925 | 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 | 926 | proof cases | 
| 61808 | 927 | assume "u x = 0" thus ?thesis using \<open>x \<in> space M\<close> f(3) by simp | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 928 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 929 | assume "u x \<noteq> 0" | 
| 61808 | 930 | with \<open>a < 1\<close> u_range[OF \<open>x \<in> space M\<close>] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 931 | have "a * u x < 1 * u x" | 
| 43920 | 932 | by (intro ereal_mult_strict_right_mono) (auto simp: image_iff) | 
| 46884 | 933 | 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 | 934 | finally obtain i where "a * u x < f i x" unfolding SUP_def | 
| 56166 | 935 | 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 | 936 | hence "a * u x \<le> f i x" by auto | 
| 61808 | 937 | thus ?thesis using \<open>x \<in> space M\<close> by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 938 | qed | 
| 40859 | 939 | qed | 
| 47694 | 940 | 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 | 941 | qed | 
| 38656 | 942 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 943 | have "integral\<^sup>S M u = (SUP i. integral\<^sup>S M (?uB i))" | 
| 61808 | 944 | unfolding simple_integral_indicator[OF B \<open>simple_function M u\<close>] | 
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 945 | proof (subst SUP_ereal_setsum, safe) | 
| 38656 | 946 | fix x n assume "x \<in> space M" | 
| 47694 | 947 | 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)" | 
| 948 | using B_mono B_u by (auto intro!: emeasure_mono ereal_mult_left_mono incseq_SucI simp: ereal_zero_le_0_iff) | |
| 38656 | 949 | next | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 950 | 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 | 951 | using measure_conv u_range B_u unfolding simple_integral_def | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59426diff
changeset | 952 | by (auto intro!: setsum.cong SUP_ereal_mult_left [symmetric]) | 
| 38656 | 953 | qed | 
| 954 | moreover | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 955 | have "a * (SUP i. integral\<^sup>S M (?uB i)) \<le> ?S" | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59426diff
changeset | 956 | apply (subst SUP_ereal_mult_left [symmetric]) | 
| 38705 | 957 | proof (safe intro!: SUP_mono bexI) | 
| 38656 | 958 | fix i | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 959 | have "a * integral\<^sup>S M (?uB i) = (\<integral>\<^sup>Sx. a * ?uB i x \<partial>M)" | 
| 61808 | 960 | using B \<open>simple_function M u\<close> u_range | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 961 | by (subst simple_integral_mult) (auto split: split_indicator) | 
| 56996 | 962 | also have "\<dots> \<le> integral\<^sup>N M (f i)" | 
| 38656 | 963 | proof - | 
| 61808 | 964 | have *: "simple_function M (\<lambda>x. a * ?uB i x)" using B \<open>0 < a\<close> u(1) by auto | 
| 965 | show ?thesis using f(3) * u_range \<open>0 < a\<close> | |
| 56996 | 966 | by (subst nn_integral_eq_simple_integral[symmetric]) | 
| 967 | (auto intro!: nn_integral_mono split: split_indicator) | |
| 38656 | 968 | qed | 
| 56996 | 969 | finally show "a * integral\<^sup>S M (?uB i) \<le> integral\<^sup>N M (f i)" | 
| 38656 | 970 | by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 971 | next | 
| 61808 | 972 | fix i show "0 \<le> \<integral>\<^sup>S x. ?uB i x \<partial>M" using B \<open>0 < a\<close> u(1) u_range | 
| 56996 | 973 | by (intro simple_integral_nonneg) (auto split: split_indicator) | 
| 61808 | 974 | qed (insert \<open>0 < a\<close>, auto) | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 975 | ultimately show "a * integral\<^sup>S M u \<le> ?S" by simp | 
| 35582 | 976 | qed | 
| 977 | ||
| 56996 | 978 | lemma incseq_nn_integral: | 
| 979 | 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 | 980 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 981 | 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 | 982 | 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 | 983 | then show ?thesis | 
| 56996 | 984 | 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 | 985 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 986 | |
| 59000 | 987 | lemma nn_integral_max_0: "(\<integral>\<^sup>+x. max 0 (f x) \<partial>M) = integral\<^sup>N M f" | 
| 988 | by (simp add: le_fun_def nn_integral_def) | |
| 989 | ||
| 61808 | 990 | text \<open>Beppo-Levi monotone convergence theorem\<close> | 
| 56996 | 991 | lemma nn_integral_monotone_convergence_SUP: | 
| 59000 | 992 | assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" | 
| 56996 | 993 | 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 | 994 | proof (rule antisym) | 
| 56996 | 995 | show "(SUP j. integral\<^sup>N M (f j)) \<le> (\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M)" | 
| 996 | by (auto intro!: SUP_least SUP_upper nn_integral_mono) | |
| 38656 | 997 | next | 
| 59000 | 998 | have f': "incseq (\<lambda>i x. max 0 (f i x))" | 
| 999 | using f by (auto simp: incseq_def le_fun_def not_le split: split_max) | |
| 1000 | (blast intro: order_trans less_imp_le) | |
| 1001 | have "(\<integral>\<^sup>+ x. max 0 (SUP i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (SUP i. max 0 (f i x)) \<partial>M)" | |
| 1002 | unfolding sup_max[symmetric] Complete_Lattices.SUP_sup_distrib[symmetric] by simp | |
| 1003 | also have "\<dots> \<le> (SUP i. (\<integral>\<^sup>+ x. max 0 (f i x) \<partial>M))" | |
| 1004 | unfolding nn_integral_def_finite[of _ "\<lambda>x. SUP i. max 0 (f i x)"] | |
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 1005 | proof (safe intro!: SUP_least) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1006 | fix g assume g: "simple_function M g" | 
| 59000 | 1007 |       and *: "g \<le> max 0 \<circ> (\<lambda>x. SUP i. max 0 (f i x))" "range g \<subseteq> {0..<\<infinity>}"
 | 
| 1008 |     then have "\<And>x. 0 \<le> (SUP i. max 0 (f i x))" and g': "g`space M \<subseteq> {0..<\<infinity>}"
 | |
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 1009 | using f by (auto intro!: SUP_upper2) | 
| 59000 | 1010 | with * show "integral\<^sup>S M g \<le> (SUP j. \<integral>\<^sup>+x. max 0 (f j x) \<partial>M)" | 
| 1011 | by (intro nn_integral_SUP_approx[OF f' _ _ g _ g']) | |
| 1012 | (auto simp: le_fun_def max_def intro!: measurable_If f borel_measurable_le) | |
| 35582 | 1013 | qed | 
| 59000 | 1014 | finally show "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) \<le> (SUP j. integral\<^sup>N M (f j))" | 
| 1015 | unfolding nn_integral_max_0 . | |
| 35582 | 1016 | qed | 
| 1017 | ||
| 56996 | 1018 | lemma nn_integral_monotone_convergence_SUP_AE: | 
| 47694 | 1019 | 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 | 1020 | shows "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (SUP i. integral\<^sup>N M (f i))" | 
| 40859 | 1021 | proof - | 
| 47694 | 1022 | 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 | 1023 | by (simp add: AE_all_countable) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1024 | from this[THEN AE_E] guess N . note N = this | 
| 46731 | 1025 | let ?f = "\<lambda>i x. if x \<in> space M - N then f i x else 0" | 
| 47694 | 1026 | 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 | 1027 | then have "(\<integral>\<^sup>+ x. (SUP i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (SUP i. ?f i x) \<partial>M)" | 
| 56996 | 1028 | 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 | 1029 | also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. ?f i x \<partial>M))" | 
| 56996 | 1030 | 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 | 1031 | 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 | 1032 |     { fix i show "(\<lambda>x. if x \<in> space M - N then f i x else 0) \<in> borel_measurable M"
 | 
| 59000 | 1033 | using f N(3) by (intro measurable_If_set) auto } | 
| 40859 | 1034 | qed | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1035 | also have "\<dots> = (SUP i. (\<integral>\<^sup>+ x. f i x \<partial>M))" | 
| 56996 | 1036 | 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 | 1037 | finally show ?thesis . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1038 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1039 | |
| 56996 | 1040 | lemma nn_integral_monotone_convergence_SUP_AE_incseq: | 
| 47694 | 1041 | 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 | 1042 | 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 | 1043 | using f[unfolded incseq_Suc_iff le_fun_def] | 
| 56996 | 1044 | 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 | 1045 | auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1046 | |
| 56996 | 1047 | lemma nn_integral_monotone_convergence_simple: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1048 | 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 | 1049 | shows "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)" | 
| 56996 | 1050 | using assms unfolding nn_integral_monotone_convergence_SUP[OF f(1) | 
| 59000 | 1051 | f(3)[THEN borel_measurable_simple_function]] | 
| 56996 | 1052 | 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 | 1053 | |
| 56996 | 1054 | lemma nn_integral_cong_pos: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1055 | assumes "\<And>x. x \<in> space M \<Longrightarrow> f x \<le> 0 \<and> g x \<le> 0 \<or> f x = g x" | 
| 56996 | 1056 | 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 | 1057 | proof - | 
| 56996 | 1058 | have "integral\<^sup>N M (\<lambda>x. max 0 (f x)) = integral\<^sup>N M (\<lambda>x. max 0 (g x))" | 
| 1059 | proof (intro nn_integral_cong) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1060 | fix x assume "x \<in> space M" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1061 | 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 | 1062 | by (auto split: split_max) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1063 | qed | 
| 56996 | 1064 | then show ?thesis by (simp add: nn_integral_max_0) | 
| 40859 | 1065 | qed | 
| 1066 | ||
| 47694 | 1067 | lemma SUP_simple_integral_sequences: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1068 | 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 | 1069 | and g: "incseq g" "\<And>i x. 0 \<le> g i x" "\<And>i. simple_function M (g i)" | 
| 47694 | 1070 | 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 | 1071 | 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 | 1072 | (is "SUPREMUM _ ?F = SUPREMUM _ ?G") | 
| 38656 | 1073 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1074 | have "(SUP i. integral\<^sup>S M (f i)) = (\<integral>\<^sup>+x. (SUP i. f i x) \<partial>M)" | 
| 56996 | 1075 | 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 | 1076 | also have "\<dots> = (\<integral>\<^sup>+x. (SUP i. g i x) \<partial>M)" | 
| 56996 | 1077 | unfolding eq[THEN nn_integral_cong_AE] .. | 
| 38656 | 1078 | also have "\<dots> = (SUP i. ?G i)" | 
| 56996 | 1079 | 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 | 1080 | finally show ?thesis by simp | 
| 38656 | 1081 | qed | 
| 1082 | ||
| 56996 | 1083 | 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 | 1084 | "0 \<le> c \<Longrightarrow> (\<integral>\<^sup>+ x. c \<partial>M) = c * (emeasure M) (space M)" | 
| 56996 | 1085 | by (subst nn_integral_eq_simple_integral) auto | 
| 38656 | 1086 | |
| 60175 | 1087 | lemma nn_integral_const_nonpos: "c \<le> 0 \<Longrightarrow> nn_integral M (\<lambda>x. c) = 0" | 
| 1088 | using nn_integral_max_0[of M "\<lambda>x. c"] by (simp add: max_def split: split_if_asm) | |
| 1089 | ||
| 56996 | 1090 | lemma nn_integral_linear: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1091 | 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 | 1092 | and g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x" | 
| 56996 | 1093 | shows "(\<integral>\<^sup>+ x. a * f x + g x \<partial>M) = a * integral\<^sup>N M f + integral\<^sup>N M g" | 
| 1094 | (is "integral\<^sup>N M ?L = _") | |
| 35582 | 1095 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1096 | from borel_measurable_implies_simple_function_sequence'[OF f(1)] guess u . | 
| 56996 | 1097 | 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 | 1098 | from borel_measurable_implies_simple_function_sequence'[OF g(1)] guess v . | 
| 56996 | 1099 | note v = nn_integral_monotone_convergence_simple[OF this(2,5,1)] this | 
| 46731 | 1100 | let ?L' = "\<lambda>i x. a * u i x + v i x" | 
| 38656 | 1101 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1102 | have "?L \<in> borel_measurable M" using assms by auto | 
| 38656 | 1103 | from borel_measurable_implies_simple_function_sequence'[OF this] guess l . | 
| 56996 | 1104 | 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 | 1105 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1106 | have inc: "incseq (\<lambda>i. a * integral\<^sup>S M (u i))" "incseq (\<lambda>i. integral\<^sup>S M (v i))" | 
| 61808 | 1107 | using u v \<open>0 \<le> a\<close> | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1108 | by (auto simp: incseq_Suc_iff le_fun_def | 
| 43920 | 1109 | 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 | 1110 | 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)" | 
| 61808 | 1111 | using u v \<open>0 \<le> a\<close> 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 | 1112 |   { 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 | 1113 | by (auto split: split_if_asm) } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1114 | note not_MInf = this | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1115 | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1116 | 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 | 1117 | 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 | 1118 | show "incseq ?L'" "\<And>i x. 0 \<le> ?L' i x" "\<And>i. simple_function M (?L' i)" | 
| 61808 | 1119 | using u v \<open>0 \<le> a\<close> unfolding incseq_Suc_iff le_fun_def | 
| 56537 | 1120 | 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 | 1121 |     { fix x
 | 
| 61808 | 1122 |       { fix i have "a * u i x \<noteq> -\<infinity>" "v i x \<noteq> -\<infinity>" "u i x \<noteq> -\<infinity>" using \<open>0 \<le> a\<close> u(6)[of i x] v(6)[of i x]
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1123 | by auto } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1124 | then have "(SUP i. a * u i x + v i x) = a * (SUP i. u i x) + (SUP i. v i x)" | 
| 61808 | 1125 | using \<open>0 \<le> a\<close> u(3) v(3) u(6)[of _ x] v(6)[of _ x] | 
| 1126 | by (subst SUP_ereal_mult_left [symmetric, OF _ u(6) \<open>0 \<le> a\<close>]) | |
| 56212 
3253aaf73a01
consolidated theorem names containing INFI and SUPR: have INF and SUP instead uniformly
 haftmann parents: 
56193diff
changeset | 1127 | (auto intro!: SUP_ereal_add | 
| 56537 | 1128 | simp: incseq_Suc_iff le_fun_def add_mono ereal_mult_left_mono) } | 
| 47694 | 1129 | then show "AE x in M. (SUP i. l i x) = (SUP i. ?L' i x)" | 
| 61808 | 1130 | unfolding l(5) using \<open>0 \<le> a\<close> u(5) v(5) l(5) f(2) g(2) | 
| 56537 | 1131 | by (intro AE_I2) (auto split: split_max) | 
| 38656 | 1132 | qed | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1133 | also have "\<dots> = (SUP i. a * integral\<^sup>S M (u i) + integral\<^sup>S M (v i))" | 
| 61808 | 1134 | using u(2, 6) v(2, 6) \<open>0 \<le> a\<close> 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 | 1135 | 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 | 1136 | 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 | 1137 | unfolding l(1)[symmetric] u(1)[symmetric] v(1)[symmetric] | 
| 61808 | 1138 | apply (subst SUP_ereal_mult_left [symmetric, OF _ pos(1) \<open>0 \<le> a\<close>]) | 
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59426diff
changeset | 1139 | apply simp | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59426diff
changeset | 1140 | apply (subst SUP_ereal_add [symmetric, OF inc not_MInf]) | 
| 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59426diff
changeset | 1141 | . | 
| 56996 | 1142 | then show ?thesis by (simp add: nn_integral_max_0) | 
| 38656 | 1143 | qed | 
| 1144 | ||
| 56996 | 1145 | lemma nn_integral_cmult: | 
| 49775 
970964690b3d
remove some unneeded positivity assumptions; generalize some assumptions to AE; tuned proofs
 hoelzl parents: 
47761diff
changeset | 1146 | assumes f: "f \<in> borel_measurable M" "0 \<le> c" | 
| 56996 | 1147 | 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 | 1148 | proof - | 
| 61808 | 1149 | have [simp]: "\<And>x. c * max 0 (f x) = max 0 (c * f x)" using \<open>0 \<le> c\<close> | 
| 43920 | 1150 | 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 | 1151 | have "(\<integral>\<^sup>+ x. c * f x \<partial>M) = (\<integral>\<^sup>+ x. c * max 0 (f x) \<partial>M)" | 
| 56996 | 1152 | 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 | 1153 | then show ?thesis | 
| 61808 | 1154 | using nn_integral_linear[OF _ _ \<open>0 \<le> c\<close>, of "\<lambda>x. max 0 (f x)" _ "\<lambda>x. 0"] f | 
| 56996 | 1155 | 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 | 1156 | qed | 
| 38656 | 1157 | |
| 56996 | 1158 | lemma nn_integral_multc: | 
| 49775 
970964690b3d
remove some unneeded positivity assumptions; generalize some assumptions to AE; tuned proofs
 hoelzl parents: 
47761diff
changeset | 1159 | assumes "f \<in> borel_measurable M" "0 \<le> c" | 
| 56996 | 1160 | 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 | 1161 | unfolding mult.commute[of _ c] nn_integral_cmult[OF assms] by simp | 
| 41096 | 1162 | |
| 59000 | 1163 | lemma nn_integral_divide: | 
| 61632 | 1164 | "\<lbrakk> 0 \<le> c; f \<in> borel_measurable M \<rbrakk> \<Longrightarrow> (\<integral>\<^sup>+ x. f x / c \<partial>M) = (\<integral>\<^sup>+ x. f x \<partial>M) / c" | 
| 1165 | by(simp add: divide_ereal_def nn_integral_multc inverse_ereal_ge0I) | |
| 59000 | 1166 | |
| 56996 | 1167 | 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 | 1168 | "A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. indicator A x\<partial>M) = (emeasure M) A" | 
| 56996 | 1169 | 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 | 1170 | (auto simp: simple_integral_indicator) | 
| 38656 | 1171 | |
| 56996 | 1172 | 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 | 1173 | "0 \<le> c \<Longrightarrow> A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x. c * indicator A x \<partial>M) = c * (emeasure M) A" | 
| 56996 | 1174 | by (subst nn_integral_eq_simple_integral) | 
| 41544 | 1175 | (auto simp: simple_function_indicator simple_integral_indicator) | 
| 38656 | 1176 | |
| 56996 | 1177 | lemma nn_integral_indicator': | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1178 | 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 | 1179 | 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 | 1180 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1181 | have "(\<integral>\<^sup>+ x. indicator A x \<partial>M) = (\<integral>\<^sup>+ x. indicator (A \<inter> space M) x \<partial>M)" | 
| 56996 | 1182 | 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 | 1183 | 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 | 1184 | by simp | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1185 | finally show ?thesis . | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1186 | qed | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1187 | |
| 56996 | 1188 | lemma nn_integral_add: | 
| 47694 | 1189 | assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" | 
| 1190 | and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" | |
| 56996 | 1191 | 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 | 1192 | proof - | 
| 47694 | 1193 | have ae: "AE x in M. max 0 (f x) + max 0 (g x) = max 0 (f x + g x)" | 
| 56537 | 1194 | 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 | 1195 | have "(\<integral>\<^sup>+ x. f x + g x \<partial>M) = (\<integral>\<^sup>+ x. max 0 (f x + g x) \<partial>M)" | 
| 56996 | 1196 | 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 | 1197 | also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) + max 0 (g x) \<partial>M)" | 
| 56996 | 1198 | 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 | 1199 | also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) \<partial>M) + (\<integral>\<^sup>+ x. max 0 (g x) \<partial>M)" | 
| 56996 | 1200 | 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 | 1201 | by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1202 | finally show ?thesis | 
| 56996 | 1203 | 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 | 1204 | qed | 
| 38656 | 1205 | |
| 56996 | 1206 | lemma nn_integral_setsum: | 
| 47694 | 1207 | 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 | 1208 | shows "(\<integral>\<^sup>+ x. (\<Sum>i\<in>P. f i x) \<partial>M) = (\<Sum>i\<in>P. integral\<^sup>N M (f i))" | 
| 38656 | 1209 | proof cases | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1210 | assume f: "finite P" | 
| 47694 | 1211 | 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 | 1212 | from f this assms(1) show ?thesis | 
| 38656 | 1213 | proof induct | 
| 1214 | case (insert i P) | |
| 47694 | 1215 | then have "f i \<in> borel_measurable M" "AE x in M. 0 \<le> f i x" | 
| 1216 | "(\<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 | 1217 | by (auto intro!: setsum_nonneg) | 
| 56996 | 1218 | from nn_integral_add[OF this] | 
| 38656 | 1219 | show ?case using insert by auto | 
| 1220 | qed simp | |
| 1221 | qed simp | |
| 1222 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1223 | 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 | 1224 | 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 | 1225 | 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 | 1226 |   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 | 1227 | 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 | 1228 | proof cases | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1229 |   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 | 1230 | 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 | 1231 | 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 | 1232 | 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 | 1233 | next | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1234 |   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 | 1235 | 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 | 1236 | by (subst Max_less_iff) (auto simp: Max_ge_iff) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1237 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1238 |   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 | 1239 | 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 | 1240 | 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 | 1241 |     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 | 1242 | 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 | 1243 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1244 | 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 | 1245 | 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 | 1246 | 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 | 1247 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 1248 | |
| 56996 | 1249 | lemma nn_integral_Markov_inequality: | 
| 49775 
970964690b3d
remove some unneeded positivity assumptions; generalize some assumptions to AE; tuned proofs
 hoelzl parents: 
47761diff
changeset | 1250 | 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 | 1251 |   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 | 1252 | (is "(emeasure M) ?A \<le> _ * ?PI") | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1253 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1254 | have "?A \<in> sets M" | 
| 61808 | 1255 | using \<open>A \<in> sets M\<close> u by auto | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1256 | hence "(emeasure M) ?A = (\<integral>\<^sup>+ x. indicator ?A x \<partial>M)" | 
| 56996 | 1257 | 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 | 1258 | also have "\<dots> \<le> (\<integral>\<^sup>+ x. c * (u x * indicator A x) \<partial>M)" using u c | 
| 56996 | 1259 | by (auto intro!: nn_integral_mono_AE | 
| 43920 | 1260 | 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 | 1261 | 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 | 1262 | using assms | 
| 56996 | 1263 | 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 | 1264 | finally show ?thesis . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1265 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1266 | |
| 56996 | 1267 | lemma nn_integral_noteq_infinite: | 
| 47694 | 1268 | assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" | 
| 56996 | 1269 | and "integral\<^sup>N M g \<noteq> \<infinity>" | 
| 47694 | 1270 | 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 | 1271 | proof (rule ccontr) | 
| 47694 | 1272 | assume c: "\<not> (AE x in M. g x \<noteq> \<infinity>)" | 
| 1273 |   have "(emeasure M) {x\<in>space M. g x = \<infinity>} \<noteq> 0"
 | |
| 1274 | using c g by (auto simp add: AE_iff_null) | |
| 1275 |   moreover have "0 \<le> (emeasure M) {x\<in>space M. g x = \<infinity>}" using g by (auto intro: measurable_sets)
 | |
| 1276 |   ultimately have "0 < (emeasure M) {x\<in>space M. g x = \<infinity>}" by auto
 | |
| 1277 |   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 | 1278 |   also have "\<dots> \<le> (\<integral>\<^sup>+x. \<infinity> * indicator {x\<in>space M. g x = \<infinity>} x \<partial>M)"
 | 
| 56996 | 1279 | using g by (subst nn_integral_cmult_indicator) auto | 
| 1280 | also have "\<dots> \<le> integral\<^sup>N M g" | |
| 1281 | using assms by (auto intro!: nn_integral_mono_AE simp: indicator_def) | |
| 61808 | 1282 | finally show False using \<open>integral\<^sup>N M g \<noteq> \<infinity>\<close> by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1283 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1284 | |
| 56996 | 1285 | lemma nn_integral_PInf: | 
| 56949 | 1286 | assumes f: "f \<in> borel_measurable M" | 
| 56996 | 1287 | and not_Inf: "integral\<^sup>N M f \<noteq> \<infinity>" | 
| 56949 | 1288 |   shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
 | 
| 1289 | proof - | |
| 1290 |   have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) = (\<integral>\<^sup>+ x. \<infinity> * indicator (f -` {\<infinity>} \<inter> space M) x \<partial>M)"
 | |
| 56996 | 1291 | using f by (subst nn_integral_cmult_indicator) (auto simp: measurable_sets) | 
| 1292 | also have "\<dots> \<le> integral\<^sup>N M (\<lambda>x. max 0 (f x))" | |
| 1293 | by (auto intro!: nn_integral_mono simp: indicator_def max_def) | |
| 1294 |   finally have "\<infinity> * (emeasure M) (f -` {\<infinity>} \<inter> space M) \<le> integral\<^sup>N M f"
 | |
| 1295 | by (simp add: nn_integral_max_0) | |
| 56949 | 1296 |   moreover have "0 \<le> (emeasure M) (f -` {\<infinity>} \<inter> space M)"
 | 
| 1297 | by (rule emeasure_nonneg) | |
| 1298 | ultimately show ?thesis | |
| 1299 | using assms by (auto split: split_if_asm) | |
| 1300 | qed | |
| 1301 | ||
| 56996 | 1302 | lemma nn_integral_PInf_AE: | 
| 1303 | assumes "f \<in> borel_measurable M" "integral\<^sup>N M f \<noteq> \<infinity>" shows "AE x in M. f x \<noteq> \<infinity>" | |
| 56949 | 1304 | proof (rule AE_I) | 
| 1305 |   show "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
 | |
| 56996 | 1306 | by (rule nn_integral_PInf[OF assms]) | 
| 56949 | 1307 |   show "f -` {\<infinity>} \<inter> space M \<in> sets M"
 | 
| 1308 | using assms by (auto intro: borel_measurable_vimage) | |
| 1309 | qed auto | |
| 1310 | ||
| 1311 | lemma simple_integral_PInf: | |
| 1312 | assumes "simple_function M f" "\<And>x. 0 \<le> f x" | |
| 1313 | and "integral\<^sup>S M f \<noteq> \<infinity>" | |
| 1314 |   shows "(emeasure M) (f -` {\<infinity>} \<inter> space M) = 0"
 | |
| 56996 | 1315 | proof (rule nn_integral_PInf) | 
| 56949 | 1316 | show "f \<in> borel_measurable M" using assms by (auto intro: borel_measurable_simple_function) | 
| 56996 | 1317 | show "integral\<^sup>N M f \<noteq> \<infinity>" | 
| 1318 | using assms by (simp add: nn_integral_eq_simple_integral) | |
| 56949 | 1319 | qed | 
| 1320 | ||
| 56996 | 1321 | lemma nn_integral_diff: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1322 | assumes f: "f \<in> borel_measurable M" | 
| 47694 | 1323 | and g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" | 
| 56996 | 1324 | and fin: "integral\<^sup>N M g \<noteq> \<infinity>" | 
| 47694 | 1325 | and mono: "AE x in M. g x \<le> f x" | 
| 56996 | 1326 | shows "(\<integral>\<^sup>+ x. f x - g x \<partial>M) = integral\<^sup>N M f - integral\<^sup>N M g" | 
| 38656 | 1327 | proof - | 
| 47694 | 1328 | have diff: "(\<lambda>x. f x - g x) \<in> borel_measurable M" "AE x in M. 0 \<le> f x - g x" | 
| 43920 | 1329 | using assms by (auto intro: ereal_diff_positive) | 
| 47694 | 1330 | have pos_f: "AE x in M. 0 \<le> f x" using mono g by auto | 
| 43920 | 1331 |   { fix a b :: ereal assume "0 \<le> a" "a \<noteq> \<infinity>" "0 \<le> b" "a \<le> b" then have "b - a + a = b"
 | 
| 1332 | 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 | 1333 | note * = this | 
| 47694 | 1334 | then have "AE x in M. f x = f x - g x + g x" | 
| 56996 | 1335 | using mono nn_integral_noteq_infinite[OF g fin] assms by auto | 
| 1336 | then have **: "integral\<^sup>N M f = (\<integral>\<^sup>+x. f x - g x \<partial>M) + integral\<^sup>N M g" | |
| 1337 | unfolding nn_integral_add[OF diff g, symmetric] | |
| 1338 | by (rule nn_integral_cong_AE) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1339 | show ?thesis unfolding ** | 
| 56996 | 1340 | using fin nn_integral_nonneg[of M g] | 
| 1341 | by (cases rule: ereal2_cases[of "\<integral>\<^sup>+ x. f x - g x \<partial>M" "integral\<^sup>N M g"]) auto | |
| 38656 | 1342 | qed | 
| 1343 | ||
| 56996 | 1344 | lemma nn_integral_suminf: | 
| 47694 | 1345 | assumes f: "\<And>i. f i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> f i x" | 
| 56996 | 1346 | shows "(\<integral>\<^sup>+ x. (\<Sum>i. f i x) \<partial>M) = (\<Sum>i. integral\<^sup>N M (f i))" | 
| 38656 | 1347 | proof - | 
| 47694 | 1348 | 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 | 1349 | using assms by (auto simp: AE_all_countable) | 
| 56996 | 1350 | have "(\<Sum>i. integral\<^sup>N M (f i)) = (SUP n. \<Sum>i<n. integral\<^sup>N M (f i))" | 
| 1351 | 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 | 1352 | also have "\<dots> = (SUP n. \<integral>\<^sup>+x. (\<Sum>i<n. f i x) \<partial>M)" | 
| 56996 | 1353 | 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 | 1354 | also have "\<dots> = \<integral>\<^sup>+x. (SUP n. \<Sum>i<n. f i x) \<partial>M" using f all_pos | 
| 56996 | 1355 | 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 | 1356 | (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 | 1357 | also have "\<dots> = \<integral>\<^sup>+x. (\<Sum>i. f i x) \<partial>M" using all_pos | 
| 56996 | 1358 | 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 | 1359 | finally show ?thesis by simp | 
| 38656 | 1360 | qed | 
| 1361 | ||
| 56996 | 1362 | 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 | 1363 | 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 | 1364 | 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 | 1365 | 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 | 1366 | proof - | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1367 | have "(\<integral>\<^sup>+x. g x \<partial>M) \<le> (\<integral>\<^sup>+x. c * f x \<partial>M)" | 
| 56996 | 1368 | 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 | 1369 | also have "(\<integral>\<^sup>+x. c * f x \<partial>M) < \<infinity>" | 
| 56996 | 1370 | 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 | 1371 | 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 | 1372 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1373 | |
| 61808 | 1374 | text \<open>Fatou's lemma: convergence theorem on limes inferior\<close> | 
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1375 | |
| 56996 | 1376 | lemma nn_integral_liminf: | 
| 43920 | 1377 | fixes u :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 47694 | 1378 | assumes u: "\<And>i. u i \<in> borel_measurable M" "\<And>i. AE x in M. 0 \<le> u i x" | 
| 56996 | 1379 | shows "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))" | 
| 38656 | 1380 | proof - | 
| 47694 | 1381 | 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 | 1382 | 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 | 1383 |     (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 | 1384 | unfolding liminf_SUP_INF using pos u | 
| 56996 | 1385 | by (intro nn_integral_monotone_convergence_SUP_AE) | 
| 44937 
22c0857b8aab
removed further legacy rules from Complete_Lattices
 hoelzl parents: 
44928diff
changeset | 1386 | (elim AE_mp, auto intro!: AE_I2 intro: INF_greatest INF_superset_mono) | 
| 56996 | 1387 | 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 | 1388 | unfolding liminf_SUP_INF | 
| 56996 | 1389 | by (auto intro!: SUP_mono exI INF_greatest nn_integral_mono INF_lower) | 
| 38656 | 1390 | finally show ?thesis . | 
| 35582 | 1391 | qed | 
| 1392 | ||
| 56993 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1393 | 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 | 1394 | "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 | 1395 | 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 | 1396 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1397 | 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 | 1398 | "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 | 1399 | 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 | 1400 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1401 | 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 | 1402 | 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 | 1403 | 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 | 1404 | 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 | 1405 | |
| 56996 | 1406 | 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 | 1407 | 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 | 1408 | 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 | 1409 | 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 | 1410 | 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 | 1411 | proof - | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1412 | 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 | 1413 | 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 | 1414 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1415 | 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 | 1416 | by auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1417 | |
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1418 | 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 | 1419 | 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 | 1420 | 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 | 1421 | 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 | 1422 | 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 | 1423 | 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 | 1424 | then have "(\<integral>\<^sup>+x. limsup (\<lambda>n. u n x) \<partial>M) < \<infinity>" | 
| 56996 | 1425 | 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 | 1426 | 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 | 1427 | by simp | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1428 | qed auto | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1429 | 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 | 1430 | using w_nonneg | 
| 56996 | 1431 | 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 | 1432 | (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 | 1433 | also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^sup>+x. w x - u n x \<partial>M)" | 
| 56996 | 1434 | 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 | 1435 | 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 | 1436 | 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 | 1437 | qed simp | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1438 | 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 | 1439 | 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 | 1440 | fix i have "(\<integral>\<^sup>+x. u i x \<partial>M) < \<infinity>" | 
| 56996 | 1441 | 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 | 1442 | 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 | 1443 | 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 | 1444 | 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 | 1445 | 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 | 1446 | finally show ?thesis | 
| 56996 | 1447 | 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 | 1448 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1449 | |
| 57025 | 1450 | lemma nn_integral_LIMSEQ: | 
| 1451 | assumes f: "incseq f" "\<And>i. f i \<in> borel_measurable M" "\<And>n x. 0 \<le> f n x" | |
| 1452 | and u: "\<And>x. (\<lambda>i. f i x) ----> u x" | |
| 1453 | shows "(\<lambda>n. integral\<^sup>N M (f n)) ----> integral\<^sup>N M u" | |
| 1454 | proof - | |
| 1455 | have "(\<lambda>n. integral\<^sup>N M (f n)) ----> (SUP n. integral\<^sup>N M (f n))" | |
| 1456 | using f by (intro LIMSEQ_SUP[of "\<lambda>n. integral\<^sup>N M (f n)"] incseq_nn_integral) | |
| 1457 | also have "(SUP n. integral\<^sup>N M (f n)) = integral\<^sup>N M (\<lambda>x. SUP n. f n x)" | |
| 1458 | using f by (intro nn_integral_monotone_convergence_SUP[symmetric]) | |
| 1459 | also have "integral\<^sup>N M (\<lambda>x. SUP n. f n x) = integral\<^sup>N M (\<lambda>x. u x)" | |
| 1460 | using f by (subst SUP_Lim_ereal[OF _ u]) (auto simp: incseq_def le_fun_def) | |
| 1461 | finally show ?thesis . | |
| 1462 | qed | |
| 1463 | ||
| 56996 | 1464 | 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 | 1465 | assumes [measurable]: | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1466 | "\<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 | 1467 | 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 | 1468 | 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 | 1469 | 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 | 1470 | 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 | 1471 | proof - | 
| 56996 | 1472 | have "limsup (\<lambda>n. integral\<^sup>N M (u n)) \<le> (\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M)" | 
| 1473 | 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 | 1474 | moreover have "(\<integral>\<^sup>+ x. limsup (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)" | 
| 56996 | 1475 | 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 | 1476 | moreover have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. u n x) \<partial>M) = (\<integral>\<^sup>+ x. u' x \<partial>M)" | 
| 56996 | 1477 | using u' by (intro nn_integral_cong_AE, eventually_elim) (metis tendsto_iff_Liminf_eq_Limsup sequentially_bot) | 
| 1478 | moreover have "(\<integral>\<^sup>+x. liminf (\<lambda>n. u n x) \<partial>M) \<le> liminf (\<lambda>n. integral\<^sup>N M (u n))" | |
| 1479 | by (intro nn_integral_liminf[OF _ bound(1)]) auto | |
| 1480 | 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 | 1481 | 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 | 1482 | 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 | 1483 | 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 | 1484 | qed | 
| 
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
 hoelzl parents: 
56949diff
changeset | 1485 | |
| 60175 | 1486 | lemma nn_integral_monotone_convergence_INF': | 
| 1487 | assumes f: "decseq f" and [measurable]: "\<And>i. f i \<in> borel_measurable M" | |
| 1488 | assumes "(\<integral>\<^sup>+ x. f 0 x \<partial>M) < \<infinity>" and nn: "\<And>x i. 0 \<le> f i x" | |
| 1489 | shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))" | |
| 1490 | proof (rule LIMSEQ_unique) | |
| 1491 | show "(\<lambda>i. integral\<^sup>N M (f i)) ----> (INF i. integral\<^sup>N M (f i))" | |
| 1492 | using f by (intro LIMSEQ_INF) (auto intro!: nn_integral_mono simp: decseq_def le_fun_def) | |
| 1493 | show "(\<lambda>i. integral\<^sup>N M (f i)) ----> \<integral>\<^sup>+ x. (INF i. f i x) \<partial>M" | |
| 1494 | proof (rule nn_integral_dominated_convergence) | |
| 1495 | show "(\<integral>\<^sup>+ x. f 0 x \<partial>M) < \<infinity>" "\<And>i. f i \<in> borel_measurable M" "f 0 \<in> borel_measurable M" | |
| 1496 | by fact+ | |
| 1497 | show "\<And>j. AE x in M. 0 \<le> f j x" | |
| 1498 | using nn by auto | |
| 1499 | show "\<And>j. AE x in M. f j x \<le> f 0 x" | |
| 1500 | using f by (auto simp: decseq_def le_fun_def) | |
| 1501 | show "AE x in M. (\<lambda>i. f i x) ----> (INF i. f i x)" | |
| 1502 | using f by (auto intro!: LIMSEQ_INF simp: decseq_def le_fun_def) | |
| 1503 | show "(\<lambda>x. INF i. f i x) \<in> borel_measurable M" | |
| 1504 | by auto | |
| 1505 | qed | |
| 1506 | qed | |
| 1507 | ||
| 1508 | lemma nn_integral_monotone_convergence_INF: | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1509 | fixes f :: "nat \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1510 | assumes f: "\<And>i j x. i \<le> j \<Longrightarrow> x \<in> space M \<Longrightarrow> f j x \<le> f i x" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1511 | and [measurable]: "\<And>i. f i \<in> borel_measurable M" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1512 | and fin: "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>" | 
| 60175 | 1513 | shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))" | 
| 1514 | proof - | |
| 1515 |   { fix f :: "nat \<Rightarrow> ereal" and j assume "decseq f"
 | |
| 1516 | then have "(INF i. f i) = (INF i. f (i + j))" | |
| 1517 | apply (intro INF_eq) | |
| 1518 | apply (rule_tac x="i" in bexI) | |
| 1519 | apply (auto simp: decseq_def le_fun_def) | |
| 1520 | done } | |
| 1521 | note INF_shift = this | |
| 1522 | ||
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1523 | have dec: "decseq (\<lambda>j x. max 0 (restrict (f (j + i)) (space M) x))" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1524 | using f by (intro antimonoI le_funI max.mono) (auto simp: decseq_def le_fun_def) | 
| 60175 | 1525 | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1526 | have "(\<integral>\<^sup>+ x. max 0 (INF i. f i x) \<partial>M) = (\<integral>\<^sup>+ x. (INF i. max 0 (restrict (f i) (space M) x)) \<partial>M)" | 
| 60175 | 1527 | by (intro nn_integral_cong) | 
| 1528 | (simp add: sup_ereal_def[symmetric] sup_INF del: sup_ereal_def) | |
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1529 | also have "\<dots> = (\<integral>\<^sup>+ x. (INF j. max 0 (restrict (f (j + i)) (space M) x)) \<partial>M)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1530 | using f by (intro nn_integral_cong INF_shift antimonoI le_funI max.mono) | 
| 60175 | 1531 | (auto simp: decseq_def le_fun_def) | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1532 | also have "\<dots> = (INF j. (\<integral>\<^sup>+ x. max 0 (restrict (f (j + i)) (space M) x) \<partial>M))" | 
| 60175 | 1533 | proof (rule nn_integral_monotone_convergence_INF') | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1534 | show "(\<lambda>x. max 0 (restrict (f (j + i)) (space M) x)) \<in> borel_measurable M" for j | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1535 | by (subst measurable_cong[where g="\<lambda>x. max 0 (f (j + i) x)"]) measurable | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1536 | show "(\<integral>\<^sup>+ x. max 0 (restrict (f (0 + i)) (space M) x) \<partial>M) < \<infinity>" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1537 | using fin by (simp add: nn_integral_max_0 cong: nn_integral_cong) | 
| 60175 | 1538 | qed (intro max.cobounded1 dec f)+ | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1539 | also have "\<dots> = (INF j. (\<integral>\<^sup>+ x. max 0 (restrict (f j) (space M) x) \<partial>M))" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1540 | using f by (intro INF_shift[symmetric] nn_integral_mono antimonoI le_funI max.mono) | 
| 60175 | 1541 | (auto simp: decseq_def le_fun_def) | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1542 | finally show ?thesis unfolding nn_integral_max_0 by (simp cong: nn_integral_cong) | 
| 60175 | 1543 | qed | 
| 1544 | ||
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1545 | lemma nn_integral_monotone_convergence_INF_decseq: | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1546 | assumes f: "decseq f" and *: "\<And>i. f i \<in> borel_measurable M" "(\<integral>\<^sup>+ x. f i x \<partial>M) < \<infinity>" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1547 | shows "(\<integral>\<^sup>+ x. (INF i. f i x) \<partial>M) = (INF i. integral\<^sup>N M (f i))" | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1548 | using nn_integral_monotone_convergence_INF[of M f i, OF _ *] f by (auto simp: antimono_def le_fun_def) | 
| 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1549 | |
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1550 | lemma sup_continuous_nn_integral[order_continuous_intros]: | 
| 60175 | 1551 | assumes f: "\<And>y. sup_continuous (f y)" | 
| 60614 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1552 | assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M" | 
| 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1553 | shows "sup_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))" | 
| 60175 | 1554 | unfolding sup_continuous_def | 
| 1555 | proof safe | |
| 60614 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1556 | fix C :: "nat \<Rightarrow> 'b" assume C: "incseq C" | 
| 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1557 | with sup_continuous_mono[OF f] show "(\<integral>\<^sup>+ y. f y (SUPREMUM UNIV C) \<partial>M) = (SUP i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)" | 
| 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1558 | unfolding sup_continuousD[OF f C] | 
| 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1559 | by (subst nn_integral_monotone_convergence_SUP) (auto simp: mono_def le_fun_def) | 
| 60175 | 1560 | qed | 
| 1561 | ||
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1562 | lemma inf_continuous_nn_integral[order_continuous_intros]: | 
| 60175 | 1563 | assumes f: "\<And>y. inf_continuous (f y)" | 
| 60614 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1564 | assumes [measurable]: "\<And>x. (\<lambda>y. f y x) \<in> borel_measurable M" | 
| 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1565 | assumes bnd: "\<And>x. (\<integral>\<^sup>+ y. f y x \<partial>M) \<noteq> \<infinity>" | 
| 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1566 | shows "inf_continuous (\<lambda>x. (\<integral>\<^sup>+y. f y x \<partial>M))" | 
| 60175 | 1567 | unfolding inf_continuous_def | 
| 1568 | proof safe | |
| 60614 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1569 | fix C :: "nat \<Rightarrow> 'b" assume C: "decseq C" | 
| 
e39e6881985c
generalized inf and sup_continuous; added intro rules
 hoelzl parents: 
60175diff
changeset | 1570 | then show "(\<integral>\<^sup>+ y. f y (INFIMUM UNIV C) \<partial>M) = (INF i. \<integral>\<^sup>+ y. f y (C i) \<partial>M)" | 
| 60175 | 1571 | using inf_continuous_mono[OF f] | 
| 1572 | by (auto simp add: inf_continuousD[OF f C] fun_eq_iff antimono_def mono_def le_fun_def bnd | |
| 1573 | intro!: nn_integral_monotone_convergence_INF) | |
| 1574 | qed | |
| 1575 | ||
| 56996 | 1576 | 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 | 1577 | assumes "N \<in> null_sets M" shows "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = 0" | 
| 38656 | 1578 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1579 | have "(\<integral>\<^sup>+ x. u x * indicator N x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)" | 
| 56996 | 1580 | proof (intro nn_integral_cong_AE AE_I) | 
| 40859 | 1581 |     show "{x \<in> space M. u x * indicator N x \<noteq> 0} \<subseteq> N"
 | 
| 1582 | by (auto simp: indicator_def) | |
| 47694 | 1583 | show "(emeasure M) N = 0" "N \<in> sets M" | 
| 40859 | 1584 | using assms by auto | 
| 35582 | 1585 | qed | 
| 40859 | 1586 | then show ?thesis by simp | 
| 38656 | 1587 | qed | 
| 35582 | 1588 | |
| 56996 | 1589 | lemma nn_integral_0_iff: | 
| 47694 | 1590 | assumes u: "u \<in> borel_measurable M" and pos: "AE x in M. 0 \<le> u x" | 
| 56996 | 1591 |   shows "integral\<^sup>N M u = 0 \<longleftrightarrow> emeasure M {x\<in>space M. u x \<noteq> 0} = 0"
 | 
| 47694 | 1592 | (is "_ \<longleftrightarrow> (emeasure M) ?A = 0") | 
| 35582 | 1593 | proof - | 
| 56996 | 1594 | have u_eq: "(\<integral>\<^sup>+ x. u x * indicator ?A x \<partial>M) = integral\<^sup>N M u" | 
| 1595 | by (auto intro!: nn_integral_cong simp: indicator_def) | |
| 38656 | 1596 | show ?thesis | 
| 1597 | proof | |
| 47694 | 1598 | assume "(emeasure M) ?A = 0" | 
| 56996 | 1599 | with nn_integral_null_set[of ?A M u] u | 
| 1600 | show "integral\<^sup>N M u = 0" by (simp add: u_eq null_sets_def) | |
| 38656 | 1601 | next | 
| 43920 | 1602 |     { fix r :: ereal and n :: nat assume gt_1: "1 \<le> real n * r"
 | 
| 1603 | then have "0 < real n * r" by (cases r) (auto split: split_if_asm simp: one_ereal_def) | |
| 1604 | 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 | 1605 | note gt_1 = this | 
| 56996 | 1606 | assume *: "integral\<^sup>N M u = 0" | 
| 46731 | 1607 |     let ?M = "\<lambda>n. {x \<in> space M. 1 \<le> real (n::nat) * u x}"
 | 
| 47694 | 1608 | have "0 = (SUP n. (emeasure M) (?M n \<inter> ?A))" | 
| 38656 | 1609 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1610 |       { fix n :: nat
 | 
| 56996 | 1611 | from nn_integral_Markov_inequality[OF u pos, of ?A "ereal (real n)"] | 
| 47694 | 1612 | have "(emeasure M) (?M n \<inter> ?A) \<le> 0" unfolding u_eq * using u by simp | 
| 1613 | moreover have "0 \<le> (emeasure M) (?M n \<inter> ?A)" using u by auto | |
| 1614 | ultimately have "(emeasure M) (?M n \<inter> ?A) = 0" by auto } | |
| 38656 | 1615 | thus ?thesis by simp | 
| 35582 | 1616 | qed | 
| 47694 | 1617 | also have "\<dots> = (emeasure M) (\<Union>n. ?M n \<inter> ?A)" | 
| 1618 | proof (safe intro!: SUP_emeasure_incseq) | |
| 38656 | 1619 | 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 | 1620 | using u by (auto intro!: sets.Int) | 
| 38656 | 1621 | next | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1622 |       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 | 1623 | proof (safe intro!: incseq_SucI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1624 | fix n :: nat and x | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1625 | assume *: "1 \<le> real n * u x" | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 1626 | also from gt_1[OF *] have "real n * u x \<le> real (Suc n) * u x" | 
| 61808 | 1627 | using \<open>0 \<le> u x\<close> 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 | 1628 | 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 | 1629 | qed | 
| 38656 | 1630 | qed | 
| 47694 | 1631 |     also have "\<dots> = (emeasure M) {x\<in>space M. 0 < u x}"
 | 
| 1632 | 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 | 1633 | fix x assume "0 < u x" and [simp, intro]: "x \<in> space M" | 
| 38656 | 1634 | show "x \<in> (\<Union>n. ?M n \<inter> ?A)" | 
| 1635 | proof (cases "u x") | |
| 61808 | 1636 | case (real r) with \<open>0 < u x\<close> have "0 < r" by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1637 | obtain j :: nat where "1 / r \<le> real j" using real_arch_simple .. | 
| 61808 | 1638 | hence "1 / r * r \<le> real j * r" unfolding mult_le_cancel_right using \<open>0 < r\<close> by auto | 
| 1639 | hence "1 \<le> real j * r" using real \<open>0 < r\<close> by auto | |
| 1640 | thus ?thesis using \<open>0 < r\<close> real by (auto simp: one_ereal_def) | |
| 1641 | qed (insert \<open>0 < u x\<close>, auto) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1642 | qed auto | 
| 47694 | 1643 |     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 | 1644 | moreover | 
| 47694 | 1645 | from pos have "AE x in M. \<not> (u x < 0)" by auto | 
| 1646 |     then have "(emeasure M) {x\<in>space M. u x < 0} = 0"
 | |
| 1647 | using AE_iff_null[of M] u by auto | |
| 1648 |     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}"
 | |
| 1649 | using u by (subst plus_emeasure) (auto intro!: arg_cong[where f="emeasure M"]) | |
| 1650 | ultimately show "(emeasure M) ?A = 0" by simp | |
| 35582 | 1651 | qed | 
| 1652 | qed | |
| 1653 | ||
| 56996 | 1654 | lemma nn_integral_0_iff_AE: | 
| 41705 | 1655 | assumes u: "u \<in> borel_measurable M" | 
| 56996 | 1656 | shows "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. u x \<le> 0)" | 
| 41705 | 1657 | proof - | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 1658 |   have sets: "{x\<in>space M. max 0 (u x) \<noteq> 0} \<in> sets M"
 | 
| 41705 | 1659 | using u by auto | 
| 56996 | 1660 | from nn_integral_0_iff[of "\<lambda>x. max 0 (u x)"] | 
| 1661 | have "integral\<^sup>N M u = 0 \<longleftrightarrow> (AE x in M. max 0 (u x) = 0)" | |
| 1662 | unfolding nn_integral_max_0 | |
| 47694 | 1663 | using AE_iff_null[OF sets] u by auto | 
| 1664 | 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 | 1665 | finally show ?thesis . | 
| 41705 | 1666 | qed | 
| 1667 | ||
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1668 | lemma AE_iff_nn_integral: | 
| 56996 | 1669 |   "{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"
 | 
| 1670 | 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 | 1671 | 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 | 1672 | |
| 59000 | 1673 | lemma nn_integral_less: | 
| 1674 | assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" | |
| 1675 | assumes f: "AE x in M. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>M) \<noteq> \<infinity>" | |
| 1676 | assumes ord: "AE x in M. f x \<le> g x" "\<not> (AE x in M. g x \<le> f x)" | |
| 1677 | shows "(\<integral>\<^sup>+x. f x \<partial>M) < (\<integral>\<^sup>+x. g x \<partial>M)" | |
| 1678 | proof - | |
| 1679 | have "0 < (\<integral>\<^sup>+x. g x - f x \<partial>M)" | |
| 1680 | proof (intro order_le_neq_trans nn_integral_nonneg notI) | |
| 1681 | assume "0 = (\<integral>\<^sup>+x. g x - f x \<partial>M)" | |
| 1682 | then have "AE x in M. g x - f x \<le> 0" | |
| 1683 | using nn_integral_0_iff_AE[of "\<lambda>x. g x - f x" M] by simp | |
| 1684 | with f(1) ord(1) have "AE x in M. g x \<le> f x" | |
| 1685 | by eventually_elim (auto simp: ereal_minus_le_iff) | |
| 1686 | with ord show False | |
| 1687 | by simp | |
| 1688 | qed | |
| 1689 | also have "\<dots> = (\<integral>\<^sup>+x. g x \<partial>M) - (\<integral>\<^sup>+x. f x \<partial>M)" | |
| 1690 | by (subst nn_integral_diff) (auto simp: f ord) | |
| 1691 | finally show ?thesis | |
| 1692 | by (simp add: ereal_less_minus_iff f nn_integral_nonneg) | |
| 1693 | qed | |
| 1694 | ||
| 56996 | 1695 | 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 | 1696 | "(\<integral>\<^sup>+x. a \<partial>M) = (if 0 \<le> a then a * (emeasure M) (space M) else 0)" | 
| 56996 | 1697 | by (auto intro!: nn_integral_0_iff_AE[THEN iffD2]) | 
| 42991 
3fa22920bf86
integral strong monotone; finite subadditivity for measure
 hoelzl parents: 
42950diff
changeset | 1698 | |
| 56996 | 1699 | lemma nn_integral_subalgebra: | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1700 | assumes f: "f \<in> borel_measurable N" "\<And>x. 0 \<le> f x" | 
| 47694 | 1701 | 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 | 1702 | shows "integral\<^sup>N N f = integral\<^sup>N M f" | 
| 39092 | 1703 | proof - | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1704 | 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 | 1705 | using N by (auto simp: measurable_def) | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1706 | 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 | 1707 | 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 | 1708 | 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 | 1709 | using N by auto | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1710 | from f show ?thesis | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1711 | apply induct | 
| 56996 | 1712 | apply (simp_all add: nn_integral_add nn_integral_cmult nn_integral_monotone_convergence_SUP N) | 
| 1713 | 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 | 1714 | done | 
| 39092 | 1715 | qed | 
| 1716 | ||
| 56996 | 1717 | lemma nn_integral_nat_function: | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1718 | fixes f :: "'a \<Rightarrow> nat" | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1719 | 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 | 1720 |   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 | 1721 | proof - | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1722 |   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 | 1723 | 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 | 1724 | by auto | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1725 | |
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1726 |   { fix x assume "x \<in> space M"
 | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1727 | 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 | 1728 | 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 | 1729 | 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 | 1730 | unfolding sums_ereal . | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1731 | moreover have "\<And>i. ereal (if i < f x then 1 else 0) = indicator (F i) x" | 
| 61808 | 1732 | using \<open>x \<in> space M\<close> by (simp add: one_ereal_def F_def) | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1733 | 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 | 1734 | 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 | 1735 | then have "(\<integral>\<^sup>+x. ereal (of_nat (f x)) \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>i. indicator (F i) x) \<partial>M)" | 
| 56996 | 1736 | by (simp cong: nn_integral_cong) | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1737 | also have "\<dots> = (\<Sum>i. emeasure M (F i))" | 
| 56996 | 1738 | by (simp add: nn_integral_suminf) | 
| 50097 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1739 | finally show ?thesis | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1740 | by (simp add: F_def) | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1741 | qed | 
| 
32973da2d4f7
rules for intergration: integrating nat-functions, integrals on finite measures, constant multiplication
 hoelzl parents: 
50027diff
changeset | 1742 | |
| 60175 | 1743 | lemma nn_integral_lfp: | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1744 | assumes sets[simp]: "\<And>s. sets (M s) = sets N" | 
| 60175 | 1745 | assumes f: "sup_continuous f" | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1746 | assumes g: "sup_continuous g" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1747 | assumes nonneg: "\<And>F s. 0 \<le> g F s" | 
| 60175 | 1748 | assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N" | 
| 1749 | assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s" | |
| 1750 | shows "(\<integral>\<^sup>+\<omega>. lfp f \<omega> \<partial>M s) = lfp g s" | |
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1751 | proof (subst lfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f and P="\<lambda>f. f \<in> borel_measurable N", symmetric]) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1752 | fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ereal" assume "incseq C" "\<And>i. C i \<in> borel_measurable N" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1753 | then show "(\<lambda>s. \<integral>\<^sup>+x. (SUP i. C i) x \<partial>M s) = (SUP i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1754 | unfolding SUP_apply[abs_def] | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1755 | by (subst nn_integral_monotone_convergence_SUP) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1756 | (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1757 | next | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1758 | show "\<And>x. (\<lambda>s. integral\<^sup>N (M s) bot) \<le> g x" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1759 | by (subst nn_integral_max_0[symmetric]) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1760 | (simp add: sup_ereal_def[symmetric] le_fun_def nonneg del: sup_ereal_def) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1761 | qed (auto simp add: step nonneg le_fun_def SUP_apply[abs_def] bot_fun_def intro!: meas f g) | 
| 60175 | 1762 | |
| 1763 | lemma nn_integral_gfp: | |
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1764 | assumes sets[simp]: "\<And>s. sets (M s) = sets N" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1765 | assumes f: "inf_continuous f" and g: "inf_continuous g" | 
| 60175 | 1766 | assumes meas: "\<And>F. F \<in> borel_measurable N \<Longrightarrow> f F \<in> borel_measurable N" | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1767 | assumes bound: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> (\<integral>\<^sup>+x. f F x \<partial>M s) < \<infinity>" | 
| 60175 | 1768 | assumes non_zero: "\<And>s. emeasure (M s) (space (M s)) \<noteq> 0" | 
| 1769 | assumes step: "\<And>F s. F \<in> borel_measurable N \<Longrightarrow> integral\<^sup>N (M s) (f F) = g (\<lambda>s. integral\<^sup>N (M s) F) s" | |
| 1770 | shows "(\<integral>\<^sup>+\<omega>. gfp f \<omega> \<partial>M s) = gfp g s" | |
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1771 | proof (subst gfp_transfer_bounded[where \<alpha>="\<lambda>F s. \<integral>\<^sup>+x. F x \<partial>M s" and g=g and f=f | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1772 | and P="\<lambda>F. F \<in> borel_measurable N \<and> (\<forall>s. (\<integral>\<^sup>+x. F x \<partial>M s) < \<infinity>)", symmetric]) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1773 | fix C :: "nat \<Rightarrow> 'b \<Rightarrow> ereal" assume "decseq C" "\<And>i. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1774 | then show "(\<lambda>s. \<integral>\<^sup>+x. (INF i. C i) x \<partial>M s) = (INF i. (\<lambda>s. \<integral>\<^sup>+x. C i x \<partial>M s))" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1775 | unfolding INF_apply[abs_def] | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1776 | by (subst nn_integral_monotone_convergence_INF_decseq) | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1777 | (auto simp: mono_def fun_eq_iff intro!: arg_cong2[where f=emeasure] cong: measurable_cong_sets) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1778 | next | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1779 | show "\<And>x. g x \<le> (\<lambda>s. integral\<^sup>N (M s) (f top))" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1780 | by (subst step) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1781 | (auto simp add: top_fun_def top_ereal_def less_le emeasure_nonneg non_zero le_fun_def | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1782 | cong del: if_cong intro!: monoD[OF inf_continuous_mono[OF g], THEN le_funD]) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1783 | next | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1784 | fix C assume "\<And>i::nat. C i \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (C i) < \<infinity>)" "decseq C" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1785 | with bound show "INFIMUM UNIV C \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) (INFIMUM UNIV C) < \<infinity>)" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1786 | unfolding INF_apply[abs_def] | 
| 61359 
e985b52c3eb3
cleanup projective limit of probability distributions; proved Ionescu-Tulcea; used it to prove infinite prob. distribution
 hoelzl parents: 
61169diff
changeset | 1787 | by (subst nn_integral_monotone_convergence_INF_decseq) | 
| 60636 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1788 | (auto cong: measurable_cong_sets intro!: borel_measurable_INF | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1789 | simp: INF_less_iff simp del: ereal_infty_less(1)) | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1790 | next | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1791 | show "\<And>x. x \<in> borel_measurable N \<and> (\<forall>s. integral\<^sup>N (M s) x < \<infinity>) \<Longrightarrow> | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1792 | (\<lambda>s. integral\<^sup>N (M s) (f x)) = g (\<lambda>s. integral\<^sup>N (M s) x)" | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1793 | by (subst step) auto | 
| 
ee18efe9b246
add named theorems order_continuous_intros; lfp/gfp_funpow; bounded variant for lfp/gfp transfer
 hoelzl parents: 
60614diff
changeset | 1794 | qed (insert bound, auto simp add: le_fun_def INF_apply[abs_def] top_fun_def intro!: meas f g) | 
| 60175 | 1795 | |
| 61808 | 1796 | subsection \<open>Integral under concrete measures\<close> | 
| 56994 | 1797 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1798 | lemma nn_integral_empty: | 
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1799 |   assumes "space M = {}"
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1800 | shows "nn_integral M f = 0" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1801 | proof - | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1802 | have "(\<integral>\<^sup>+ x. f x \<partial>M) = (\<integral>\<^sup>+ x. 0 \<partial>M)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1803 | by(rule nn_integral_cong)(simp add: assms) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1804 | thus ?thesis by simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1805 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 1806 | |
| 61808 | 1807 | subsubsection \<open>Distributions\<close> | 
| 47694 | 1808 | |
| 56996 | 1809 | lemma nn_integral_distr': | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1810 | assumes T: "T \<in> measurable M M'" | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1811 | and f: "f \<in> borel_measurable (distr M M' T)" "\<And>x. 0 \<le> f x" | 
| 56996 | 1812 | shows "integral\<^sup>N (distr M M' T) f = (\<integral>\<^sup>+ x. f (T x) \<partial>M)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1813 | using f | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1814 | proof induct | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1815 | case (cong f g) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1816 | with T show ?case | 
| 56996 | 1817 | apply (subst nn_integral_cong[of _ f g]) | 
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1818 | apply simp | 
| 56996 | 1819 | 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 | 1820 | apply (simp add: measurable_def Pi_iff) | 
| 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 1821 | apply simp | 
| 49797 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1822 | done | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1823 | next | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1824 | case (set A) | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1825 | 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 | 1826 | by (auto simp: indicator_def) | 
| 
28066863284c
add induction rules for simple functions and for Borel measurable functions
 hoelzl parents: 
49796diff
changeset | 1827 | from set T show ?case | 
| 56996 | 1828 | by (subst nn_integral_cong[OF eq]) | 
| 1829 | (auto simp add: emeasure_distr intro!: nn_integral_indicator[symmetric] measurable_sets) | |
| 1830 | qed (simp_all add: measurable_compose[OF T] T nn_integral_cmult nn_integral_add | |
| 1831 | nn_integral_monotone_convergence_SUP le_fun_def incseq_def) | |
| 47694 | 1832 | |
| 56996 | 1833 | lemma nn_integral_distr: | 
| 1834 | "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)" | |
| 1835 | by (subst (1 2) nn_integral_max_0[symmetric]) | |
| 1836 | (simp add: nn_integral_distr') | |
| 35692 | 1837 | |
| 61808 | 1838 | subsubsection \<open>Counting space\<close> | 
| 47694 | 1839 | |
| 1840 | lemma simple_function_count_space[simp]: | |
| 1841 | "simple_function (count_space A) f \<longleftrightarrow> finite (f ` A)" | |
| 1842 | unfolding simple_function_def by simp | |
| 1843 | ||
| 56996 | 1844 | lemma nn_integral_count_space: | 
| 47694 | 1845 |   assumes A: "finite {a\<in>A. 0 < f a}"
 | 
| 56996 | 1846 | shows "integral\<^sup>N (count_space A) f = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)" | 
| 35582 | 1847 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 1848 | 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 | 1849 |     (\<integral>\<^sup>+ x. (\<Sum>a|a\<in>A \<and> 0 < f a. f a * indicator {a} x) \<partial>count_space A)"
 | 
| 56996 | 1850 | by (auto intro!: nn_integral_cong | 
| 57418 | 1851 | 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 | 1852 |   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 | 1853 | by (subst nn_integral_setsum) | 
| 47694 | 1854 | (simp_all add: AE_count_space ereal_zero_le_0_iff less_imp_le) | 
| 1855 | also have "\<dots> = (\<Sum>a|a\<in>A \<and> 0 < f a. f a)" | |
| 57418 | 1856 | by (auto intro!: setsum.cong simp: nn_integral_cmult_indicator one_ereal_def[symmetric]) | 
| 56996 | 1857 | finally show ?thesis by (simp add: nn_integral_max_0) | 
| 47694 | 1858 | qed | 
| 1859 | ||
| 56996 | 1860 | 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 | 1861 | "finite A \<Longrightarrow> (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<Sum>a\<in>A. max 0 (f a))" | 
| 56996 | 1862 | by (subst nn_integral_max_0[symmetric]) | 
| 57418 | 1863 | (auto intro!: setsum.mono_neutral_left simp: nn_integral_count_space less_le) | 
| 47694 | 1864 | |
| 59000 | 1865 | lemma nn_integral_count_space': | 
| 1866 | assumes "finite A" "\<And>x. x \<in> B \<Longrightarrow> x \<notin> A \<Longrightarrow> f x = 0" "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x" "A \<subseteq> B" | |
| 1867 | shows "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>x\<in>A. f x)" | |
| 1868 | proof - | |
| 1869 | have "(\<integral>\<^sup>+x. f x \<partial>count_space B) = (\<Sum>a | a \<in> B \<and> 0 < f a. f a)" | |
| 1870 | using assms(2,3) | |
| 61808 | 1871 | by (intro nn_integral_count_space finite_subset[OF _ \<open>finite A\<close>]) (auto simp: less_le) | 
| 59000 | 1872 | also have "\<dots> = (\<Sum>a\<in>A. f a)" | 
| 1873 | using assms by (intro setsum.mono_neutral_cong_left) (auto simp: less_le) | |
| 1874 | finally show ?thesis . | |
| 1875 | qed | |
| 1876 | ||
| 59011 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1877 | lemma nn_integral_bij_count_space: | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1878 | assumes g: "bij_betw g A B" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1879 | shows "(\<integral>\<^sup>+x. f (g x) \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)" | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1880 | using g[THEN bij_betw_imp_funcset] | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1881 | by (subst distr_bij_count_space[OF g, symmetric]) | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1882 | (auto intro!: nn_integral_distr[symmetric]) | 
| 
4902a2fec434
add reindex rules for distr and nn_integral on count_space
 hoelzl parents: 
59002diff
changeset | 1883 | |
| 59000 | 1884 | lemma nn_integral_indicator_finite: | 
| 1885 | fixes f :: "'a \<Rightarrow> ereal" | |
| 1886 |   assumes f: "finite A" and nn: "\<And>x. x \<in> A \<Longrightarrow> 0 \<le> f x" and [measurable]: "\<And>a. a \<in> A \<Longrightarrow> {a} \<in> sets M"
 | |
| 1887 |   shows "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<Sum>x\<in>A. f x * emeasure M {x})"
 | |
| 1888 | proof - | |
| 1889 |   from f have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = (\<integral>\<^sup>+x. (\<Sum>a\<in>A. f a * indicator {a} x) \<partial>M)"
 | |
| 1890 | by (intro nn_integral_cong) (auto simp: indicator_def if_distrib[where f="\<lambda>a. x * a" for x] setsum.If_cases) | |
| 1891 |   also have "\<dots> = (\<Sum>a\<in>A. f a * emeasure M {a})"
 | |
| 1892 | using nn by (subst nn_integral_setsum) (auto simp: nn_integral_cmult_indicator) | |
| 1893 | finally show ?thesis . | |
| 1894 | qed | |
| 1895 | ||
| 57025 | 1896 | lemma nn_integral_count_space_nat: | 
| 1897 | fixes f :: "nat \<Rightarrow> ereal" | |
| 1898 | assumes nonneg: "\<And>i. 0 \<le> f i" | |
| 1899 | shows "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = (\<Sum>i. f i)" | |
| 1900 | proof - | |
| 1901 | have "(\<integral>\<^sup>+i. f i \<partial>count_space UNIV) = | |
| 1902 |     (\<integral>\<^sup>+i. (\<Sum>j. f j * indicator {j} i) \<partial>count_space UNIV)"
 | |
| 1903 | proof (intro nn_integral_cong) | |
| 1904 | fix i | |
| 1905 |     have "f i = (\<Sum>j\<in>{i}. f j * indicator {j} i)"
 | |
| 1906 | by simp | |
| 1907 |     also have "\<dots> = (\<Sum>j. f j * indicator {j} i)"
 | |
| 1908 | by (rule suminf_finite[symmetric]) auto | |
| 1909 |     finally show "f i = (\<Sum>j. f j * indicator {j} i)" .
 | |
| 1910 | qed | |
| 1911 |   also have "\<dots> = (\<Sum>j. (\<integral>\<^sup>+i. f j * indicator {j} i \<partial>count_space UNIV))"
 | |
| 1912 | by (rule nn_integral_suminf) (auto simp: nonneg) | |
| 1913 | also have "\<dots> = (\<Sum>j. f j)" | |
| 1914 | by (simp add: nonneg nn_integral_cmult_indicator one_ereal_def[symmetric]) | |
| 1915 | finally show ?thesis . | |
| 1916 | qed | |
| 1917 | ||
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1918 | lemma nn_integral_count_space_nn_integral: | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1919 | fixes f :: "'i \<Rightarrow> 'a \<Rightarrow> ereal" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1920 | assumes "countable I" and [measurable]: "\<And>i. i \<in> I \<Longrightarrow> f i \<in> borel_measurable M" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1921 | shows "(\<integral>\<^sup>+x. \<integral>\<^sup>+i. f i x \<partial>count_space I \<partial>M) = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. f i x \<partial>M \<partial>count_space I)" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1922 | proof cases | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1923 | assume "finite I" then show ?thesis | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1924 | by (simp add: nn_integral_count_space_finite nn_integral_nonneg max.absorb2 nn_integral_setsum | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1925 | nn_integral_max_0) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1926 | next | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1927 | assume "infinite I" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1928 |   then have [simp]: "I \<noteq> {}"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1929 | by auto | 
| 61808 | 1930 | note * = bij_betw_from_nat_into[OF \<open>countable I\<close> \<open>infinite I\<close>] | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1931 | have **: "\<And>f. (\<And>i. 0 \<le> f i) \<Longrightarrow> (\<integral>\<^sup>+i. f i \<partial>count_space I) = (\<Sum>n. f (from_nat_into I n))" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1932 | by (simp add: nn_integral_bij_count_space[symmetric, OF *] nn_integral_count_space_nat) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1933 | show ?thesis | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1934 | apply (subst (2) nn_integral_max_0[symmetric]) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1935 | apply (simp add: ** nn_integral_nonneg nn_integral_suminf from_nat_into) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1936 | apply (simp add: nn_integral_max_0) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1937 | done | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1938 | qed | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1939 | |
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1940 | lemma emeasure_UN_countable: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1941 | assumes sets[measurable]: "\<And>i. i \<in> I \<Longrightarrow> X i \<in> sets M" and I[simp]: "countable I" | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1942 | assumes disj: "disjoint_family_on X I" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1943 | shows "emeasure M (UNION I X) = (\<integral>\<^sup>+i. emeasure M (X i) \<partial>count_space I)" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1944 | proof - | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1945 | have eq: "\<And>x. indicator (UNION I X) x = \<integral>\<^sup>+ i. indicator (X i) x \<partial>count_space I" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 1946 | proof cases | 
| 59426 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1947 | fix x assume x: "x \<in> UNION I X" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1948 | then obtain j where j: "x \<in> X j" "j \<in> I" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1949 | by auto | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1950 |     with disj have "\<And>i. i \<in> I \<Longrightarrow> indicator (X i) x = (indicator {j} i::ereal)"
 | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1951 | by (auto simp: disjoint_family_on_def split: split_indicator) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1952 | with x j show "?thesis x" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1953 | by (simp cong: nn_integral_cong_simp) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1954 | qed (auto simp: nn_integral_0_iff_AE) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1955 | |
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1956 | note sets.countable_UN'[unfolded subset_eq, measurable] | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1957 | have "emeasure M (UNION I X) = (\<integral>\<^sup>+x. indicator (UNION I X) x \<partial>M)" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1958 | by simp | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1959 | also have "\<dots> = (\<integral>\<^sup>+i. \<integral>\<^sup>+x. indicator (X i) x \<partial>M \<partial>count_space I)" | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1960 | by (simp add: eq nn_integral_count_space_nn_integral) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1961 | finally show ?thesis | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1962 | by (simp cong: nn_integral_cong_simp) | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1963 | qed | 
| 
6fca83e88417
integral of the product of count spaces equals the integral of the count space of the product type
 hoelzl parents: 
59425diff
changeset | 1964 | |
| 57025 | 1965 | lemma emeasure_countable_singleton: | 
| 1966 |   assumes sets: "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M" and X: "countable X"
 | |
| 1967 |   shows "emeasure M X = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
 | |
| 1968 | proof - | |
| 1969 |   have "emeasure M (\<Union>i\<in>X. {i}) = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space X)"
 | |
| 1970 | using assms by (intro emeasure_UN_countable) (auto simp: disjoint_family_on_def) | |
| 1971 |   also have "(\<Union>i\<in>X. {i}) = X" by auto
 | |
| 1972 | finally show ?thesis . | |
| 1973 | qed | |
| 1974 | ||
| 1975 | lemma measure_eqI_countable: | |
| 1976 | assumes [simp]: "sets M = Pow A" "sets N = Pow A" and A: "countable A" | |
| 1977 |   assumes eq: "\<And>a. a \<in> A \<Longrightarrow> emeasure M {a} = emeasure N {a}"
 | |
| 1978 | shows "M = N" | |
| 1979 | proof (rule measure_eqI) | |
| 1980 | fix X assume "X \<in> sets M" | |
| 1981 | then have X: "X \<subseteq> A" by auto | |
| 1982 | moreover with A have "countable X" by (auto dest: countable_subset) | |
| 1983 | ultimately have | |
| 1984 |     "emeasure M X = (\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X)"
 | |
| 1985 |     "emeasure N X = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
 | |
| 1986 | by (auto intro!: emeasure_countable_singleton) | |
| 1987 |   moreover have "(\<integral>\<^sup>+a. emeasure M {a} \<partial>count_space X) = (\<integral>\<^sup>+a. emeasure N {a} \<partial>count_space X)"
 | |
| 1988 | using X by (intro nn_integral_cong eq) auto | |
| 1989 | ultimately show "emeasure M X = emeasure N X" | |
| 1990 | by simp | |
| 1991 | qed simp | |
| 1992 | ||
| 59000 | 1993 | lemma measure_eqI_countable_AE: | 
| 1994 | assumes [simp]: "sets M = UNIV" "sets N = UNIV" | |
| 1995 | assumes ae: "AE x in M. x \<in> \<Omega>" "AE x in N. x \<in> \<Omega>" and [simp]: "countable \<Omega>" | |
| 1996 |   assumes eq: "\<And>x. x \<in> \<Omega> \<Longrightarrow> emeasure M {x} = emeasure N {x}"
 | |
| 1997 | shows "M = N" | |
| 1998 | proof (rule measure_eqI) | |
| 1999 | fix A | |
| 2000 |   have "emeasure N A = emeasure N {x\<in>\<Omega>. x \<in> A}"
 | |
| 2001 | using ae by (intro emeasure_eq_AE) auto | |
| 2002 |   also have "\<dots> = (\<integral>\<^sup>+x. emeasure N {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
 | |
| 2003 | by (intro emeasure_countable_singleton) auto | |
| 2004 |   also have "\<dots> = (\<integral>\<^sup>+x. emeasure M {x} \<partial>count_space {x\<in>\<Omega>. x \<in> A})"
 | |
| 2005 | by (intro nn_integral_cong eq[symmetric]) auto | |
| 2006 |   also have "\<dots> = emeasure M {x\<in>\<Omega>. x \<in> A}"
 | |
| 2007 | by (intro emeasure_countable_singleton[symmetric]) auto | |
| 2008 | also have "\<dots> = emeasure M A" | |
| 2009 | using ae by (intro emeasure_eq_AE) auto | |
| 2010 | finally show "emeasure M A = emeasure N A" .. | |
| 2011 | qed simp | |
| 2012 | ||
| 60064 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2013 | lemma nn_integral_monotone_convergence_SUP_nat': | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2014 | fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ereal" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2015 | assumes chain: "Complete_Partial_Order.chain op \<le> (f ` Y)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2016 |   and nonempty: "Y \<noteq> {}"
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2017 | and nonneg: "\<And>i n. i \<in> Y \<Longrightarrow> f i n \<ge> 0" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2018 | shows "(\<integral>\<^sup>+ x. (SUP i:Y. f i x) \<partial>count_space UNIV) = (SUP i:Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space UNIV))" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2019 | (is "?lhs = ?rhs" is "integral\<^sup>N ?M _ = _") | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2020 | proof (rule order_class.order.antisym) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2021 | show "?rhs \<le> ?lhs" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2022 | by (auto intro!: SUP_least SUP_upper nn_integral_mono) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2023 | next | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2024 | have "\<And>x. \<exists>g. incseq g \<and> range g \<subseteq> (\<lambda>i. f i x) ` Y \<and> (SUP i:Y. f i x) = (SUP i. g i)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2025 | unfolding Sup_class.SUP_def by(rule Sup_countable_SUP[unfolded Sup_class.SUP_def])(simp add: nonempty) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2026 | then obtain g where incseq: "\<And>x. incseq (g x)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2027 | and range: "\<And>x. range (g x) \<subseteq> (\<lambda>i. f i x) ` Y" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2028 | and sup: "\<And>x. (SUP i:Y. f i x) = (SUP i. g x i)" by moura | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2029 | from incseq have incseq': "incseq (\<lambda>i x. g x i)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2030 | by(blast intro: incseq_SucI le_funI dest: incseq_SucD) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2031 | |
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2032 | have "?lhs = \<integral>\<^sup>+ x. (SUP i. g x i) \<partial>?M" by(simp add: sup) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2033 | also have "\<dots> = (SUP i. \<integral>\<^sup>+ x. g x i \<partial>?M)" using incseq' | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2034 | by(rule nn_integral_monotone_convergence_SUP) simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2035 | also have "\<dots> \<le> (SUP i:Y. \<integral>\<^sup>+ x. f i x \<partial>?M)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2036 | proof(rule SUP_least) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2037 | fix n | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2038 | have "\<And>x. \<exists>i. g x n = f i x \<and> i \<in> Y" using range by blast | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2039 | then obtain I where I: "\<And>x. g x n = f (I x) x" "\<And>x. I x \<in> Y" by moura | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2040 | |
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2041 |     { fix x
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2042 | from range[of x] obtain i where "i \<in> Y" "g x n = f i x" by auto | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2043 | hence "g x n \<ge> 0" using nonneg[of i x] by simp } | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2044 | note nonneg_g = this | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2045 | then have "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) = (\<Sum>x. g x n)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2046 | by(rule nn_integral_count_space_nat) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2047 | also have "\<dots> = (SUP m. \<Sum>x<m. g x n)" using nonneg_g | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2048 | by(rule suminf_ereal_eq_SUP) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2049 | also have "\<dots> \<le> (SUP i:Y. \<integral>\<^sup>+ x. f i x \<partial>?M)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2050 | proof(rule SUP_mono) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2051 | fix m | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2052 | show "\<exists>m'\<in>Y. (\<Sum>x<m. g x n) \<le> (\<integral>\<^sup>+ x. f m' x \<partial>?M)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2053 | proof(cases "m > 0") | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2054 | case False | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2055 | thus ?thesis using nonempty by(auto simp add: nn_integral_nonneg) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2056 | next | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2057 | case True | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2058 |         let ?Y = "I ` {..<m}"
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2059 | have "f ` ?Y \<subseteq> f ` Y" using I by auto | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2060 | with chain have chain': "Complete_Partial_Order.chain op \<le> (f ` ?Y)" by(rule chain_subset) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2061 | hence "Sup (f ` ?Y) \<in> f ` ?Y" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2062 | by(rule ccpo_class.in_chain_finite)(auto simp add: True lessThan_empty_iff) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2063 | then obtain m' where "m' < m" and m': "(SUP i:?Y. f i) = f (I m')" by auto | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2064 | have "I m' \<in> Y" using I by blast | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2065 | have "(\<Sum>x<m. g x n) \<le> (\<Sum>x<m. f (I m') x)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2066 | proof(rule setsum_mono) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2067 | fix x | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2068 |           assume "x \<in> {..<m}"
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2069 | hence "x < m" by simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2070 | have "g x n = f (I x) x" by(simp add: I) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2071 | also have "\<dots> \<le> (SUP i:?Y. f i) x" unfolding SUP_def Sup_fun_def image_image | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2072 |             using \<open>x \<in> {..<m}\<close> by(rule Sup_upper[OF imageI])
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2073 | also have "\<dots> = f (I m') x" unfolding m' by simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2074 | finally show "g x n \<le> f (I m') x" . | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2075 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2076 | also have "\<dots> \<le> (SUP m. (\<Sum>x<m. f (I m') x))" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2077 | by(rule SUP_upper) simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2078 | also have "\<dots> = (\<Sum>x. f (I m') x)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2079 | by(rule suminf_ereal_eq_SUP[symmetric])(simp add: nonneg \<open>I m' \<in> Y\<close>) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2080 | also have "\<dots> = (\<integral>\<^sup>+ x. f (I m') x \<partial>?M)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2081 | by(rule nn_integral_count_space_nat[symmetric])(simp add: nonneg \<open>I m' \<in> Y\<close>) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2082 | finally show ?thesis using \<open>I m' \<in> Y\<close> by blast | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2083 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2084 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2085 | finally show "(\<integral>\<^sup>+ x. g x n \<partial>count_space UNIV) \<le> \<dots>" . | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2086 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2087 | finally show "?lhs \<le> ?rhs" . | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2088 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2089 | |
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2090 | lemma nn_integral_monotone_convergence_SUP_nat: | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2091 | fixes f :: "'a \<Rightarrow> nat \<Rightarrow> ereal" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2092 |   assumes nonempty: "Y \<noteq> {}"
 | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2093 | and chain: "Complete_Partial_Order.chain op \<le> (f ` Y)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2094 | shows "(\<integral>\<^sup>+ x. (SUP i:Y. f i x) \<partial>count_space UNIV) = (SUP i:Y. (\<integral>\<^sup>+ x. f i x \<partial>count_space UNIV))" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2095 | (is "?lhs = ?rhs") | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2096 | proof - | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2097 | let ?f = "\<lambda>i x. max 0 (f i x)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2098 | have chain': "Complete_Partial_Order.chain op \<le> (?f ` Y)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2099 | proof(rule chainI) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2100 | fix g h | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2101 | assume "g \<in> ?f ` Y" "h \<in> ?f ` Y" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2102 | then obtain g' h' where gh: "g' \<in> Y" "h' \<in> Y" "g = ?f g'" "h = ?f h'" by blast | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2103 | hence "f g' \<in> f ` Y" "f h' \<in> f ` Y" by blast+ | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2104 | with chain have "f g' \<le> f h' \<or> f h' \<le> f g'" by(rule chainD) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2105 | thus "g \<le> h \<or> h \<le> g" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2106 | proof | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2107 | assume "f g' \<le> f h'" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2108 | hence "g \<le> h" using gh order_trans by(auto simp add: le_fun_def max_def) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2109 | thus ?thesis .. | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2110 | next | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2111 | assume "f h' \<le> f g'" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2112 | hence "h \<le> g" using gh order_trans by(auto simp add: le_fun_def max_def) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2113 | thus ?thesis .. | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2114 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2115 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2116 | have "?lhs = (\<integral>\<^sup>+ x. max 0 (SUP i:Y. f i x) \<partial>count_space UNIV)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2117 | by(simp add: nn_integral_max_0) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2118 | also have "\<dots> = (\<integral>\<^sup>+ x. (SUP i:Y. ?f i x) \<partial>count_space UNIV)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2119 | proof(rule nn_integral_cong) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2120 | fix x | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2121 | have "max 0 (SUP i:Y. f i x) \<le> (SUP i:Y. ?f i x)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2122 | proof(cases "0 \<le> (SUP i:Y. f i x)") | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2123 | case True | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2124 | have "(SUP i:Y. f i x) \<le> (SUP i:Y. ?f i x)" by(rule SUP_mono)(auto intro: rev_bexI) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2125 | with True show ?thesis by simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2126 | next | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2127 | case False | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2128 | have "0 \<le> (SUP i:Y. ?f i x)" using nonempty by(auto intro: SUP_upper2) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2129 | thus ?thesis using False by simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2130 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2131 | moreover have "\<dots> \<le> max 0 (SUP i:Y. f i x)" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2132 | proof(cases "(SUP i:Y. f i x) \<ge> 0") | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2133 | case True | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2134 | show ?thesis | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2135 | by(rule SUP_least)(auto simp add: True max_def intro: SUP_upper) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2136 | next | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2137 | case False | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2138 | hence "(SUP i:Y. f i x) \<le> 0" by simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2139 | hence less: "\<forall>i\<in>Y. f i x \<le> 0" by(simp add: SUP_le_iff) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2140 | show ?thesis by(rule SUP_least)(auto simp add: max_def less intro: SUP_upper) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2141 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2142 | ultimately show "\<dots> = (SUP i:Y. ?f i x)" by(rule order.antisym) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2143 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2144 | also have "\<dots> = (SUP i:Y. (\<integral>\<^sup>+ x. ?f i x \<partial>count_space UNIV))" | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2145 | using chain' nonempty by(rule nn_integral_monotone_convergence_SUP_nat') simp | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2146 | also have "\<dots> = ?rhs" by(simp add: nn_integral_max_0) | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2147 | finally show ?thesis . | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2148 | qed | 
| 
63124d48a2ee
add lemma about monotone convergence for countable integrals over arbitrary sequences
 Andreas Lochbihler parents: 
59779diff
changeset | 2149 | |
| 61808 | 2150 | subsubsection \<open>Measures with Restricted Space\<close> | 
| 54417 | 2151 | |
| 57137 | 2152 | lemma simple_function_iff_borel_measurable: | 
| 2153 |   fixes f :: "'a \<Rightarrow> 'x::{t2_space}"
 | |
| 2154 | shows "simple_function M f \<longleftrightarrow> finite (f ` space M) \<and> f \<in> borel_measurable M" | |
| 2155 | by (metis borel_measurable_simple_function simple_functionD(1) simple_function_borel_measurable) | |
| 2156 | ||
| 2157 | lemma simple_function_restrict_space_ereal: | |
| 2158 | fixes f :: "'a \<Rightarrow> ereal" | |
| 2159 | assumes "\<Omega> \<inter> space M \<in> sets M" | |
| 2160 | shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 2161 | proof - | |
| 2162 |   { assume "finite (f ` space (restrict_space M \<Omega>))"
 | |
| 2163 |     then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
 | |
| 2164 | then have "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)" | |
| 2165 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 2166 | moreover | |
| 2167 |   { assume "finite ((\<lambda>x. f x * indicator \<Omega> x) ` space M)"
 | |
| 2168 | then have "finite (f ` space (restrict_space M \<Omega>))" | |
| 2169 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 2170 | ultimately show ?thesis | |
| 2171 | unfolding simple_function_iff_borel_measurable | |
| 2172 | borel_measurable_restrict_space_iff_ereal[OF assms] | |
| 2173 | by auto | |
| 2174 | qed | |
| 2175 | ||
| 2176 | lemma simple_function_restrict_space: | |
| 2177 | fixes f :: "'a \<Rightarrow> 'b::real_normed_vector" | |
| 2178 | assumes "\<Omega> \<inter> space M \<in> sets M" | |
| 2179 | shows "simple_function (restrict_space M \<Omega>) f \<longleftrightarrow> simple_function M (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" | |
| 2180 | proof - | |
| 2181 |   { assume "finite (f ` space (restrict_space M \<Omega>))"
 | |
| 2182 |     then have "finite (f ` space (restrict_space M \<Omega>) \<union> {0})" by simp
 | |
| 2183 | then have "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)" | |
| 2184 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 2185 | moreover | |
| 2186 |   { assume "finite ((\<lambda>x. indicator \<Omega> x *\<^sub>R f x) ` space M)"
 | |
| 2187 | then have "finite (f ` space (restrict_space M \<Omega>))" | |
| 2188 | by (rule rev_finite_subset) (auto split: split_indicator simp: space_restrict_space) } | |
| 2189 | ultimately show ?thesis | |
| 2190 | unfolding simple_function_iff_borel_measurable | |
| 2191 | borel_measurable_restrict_space_iff[OF assms] | |
| 2192 | by auto | |
| 2193 | qed | |
| 2194 | ||
| 2195 | ||
| 2196 | lemma split_indicator_asm: "P (indicator Q x) = (\<not> (x \<in> Q \<and> \<not> P 1 \<or> x \<notin> Q \<and> \<not> P 0))" | |
| 2197 | by (auto split: split_indicator) | |
| 2198 | ||
| 2199 | lemma simple_integral_restrict_space: | |
| 2200 | assumes \<Omega>: "\<Omega> \<inter> space M \<in> sets M" "simple_function (restrict_space M \<Omega>) f" | |
| 2201 | shows "simple_integral (restrict_space M \<Omega>) f = simple_integral M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 2202 | using simple_function_restrict_space_ereal[THEN iffD1, OF \<Omega>, THEN simple_functionD(1)] | |
| 2203 | by (auto simp add: space_restrict_space emeasure_restrict_space[OF \<Omega>(1)] le_infI2 simple_integral_def | |
| 2204 | split: split_indicator split_indicator_asm | |
| 59002 
2c8b2fb54b88
cleaning up some theorem names; remove unnecessary assumptions; more complete pmf theory
 hoelzl parents: 
59000diff
changeset | 2205 | intro!: setsum.mono_neutral_cong_left ereal_right_mult_cong[OF refl] arg_cong2[where f=emeasure]) | 
| 57137 | 2206 | |
| 2207 | lemma one_not_less_zero[simp]: "\<not> 1 < (0::ereal)" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2208 | by (simp add: zero_ereal_def one_ereal_def) | 
| 57137 | 2209 | |
| 56996 | 2210 | lemma nn_integral_restrict_space: | 
| 57137 | 2211 | assumes \<Omega>[simp]: "\<Omega> \<inter> space M \<in> sets M" | 
| 2212 | shows "nn_integral (restrict_space M \<Omega>) f = nn_integral M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 2213 | proof - | |
| 2214 |   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>}}"
 | |
| 2215 | have "integral\<^sup>S ?R ` ?X ?R f = integral\<^sup>S M ` ?X M (\<lambda>x. f x * indicator \<Omega> x)" | |
| 2216 | proof (safe intro!: image_eqI) | |
| 2217 |     fix s assume s: "simple_function ?R s" "s \<le> max 0 \<circ> f" "range s \<subseteq> {0..<\<infinity>}"
 | |
| 2218 | from s show "integral\<^sup>S (restrict_space M \<Omega>) s = integral\<^sup>S M (\<lambda>x. s x * indicator \<Omega> x)" | |
| 2219 | by (intro simple_integral_restrict_space) auto | |
| 2220 | from s show "simple_function M (\<lambda>x. s x * indicator \<Omega> x)" | |
| 2221 | by (simp add: simple_function_restrict_space_ereal) | |
| 2222 | from s show "(\<lambda>x. s x * indicator \<Omega> x) \<le> max 0 \<circ> (\<lambda>x. f x * indicator \<Omega> x)" | |
| 2223 |       "\<And>x. s x * indicator \<Omega> x \<in> {0..<\<infinity>}"
 | |
| 2224 | by (auto split: split_indicator simp: le_fun_def image_subset_iff) | |
| 2225 | next | |
| 2226 |     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>}"
 | |
| 2227 | then have "simple_function M (\<lambda>x. s x * indicator (\<Omega> \<inter> space M) x)" (is ?s') | |
| 2228 | by (intro simple_function_mult simple_function_indicator) auto | |
| 2229 | also have "?s' \<longleftrightarrow> simple_function M (\<lambda>x. s x * indicator \<Omega> x)" | |
| 2230 | by (rule simple_function_cong) (auto split: split_indicator) | |
| 2231 | finally show sf: "simple_function (restrict_space M \<Omega>) s" | |
| 2232 | by (simp add: simple_function_restrict_space_ereal) | |
| 2233 | ||
| 2234 | from s have s_eq: "s = (\<lambda>x. s x * indicator \<Omega> x)" | |
| 2235 | by (auto simp add: fun_eq_iff le_fun_def image_subset_iff | |
| 2236 | split: split_indicator split_indicator_asm | |
| 2237 | intro: antisym) | |
| 2238 | ||
| 2239 | show "integral\<^sup>S M s = integral\<^sup>S (restrict_space M \<Omega>) s" | |
| 2240 | by (subst s_eq) (rule simple_integral_restrict_space[symmetric, OF \<Omega> sf]) | |
| 2241 |     show "\<And>x. s x \<in> {0..<\<infinity>}"
 | |
| 2242 | using s by (auto simp: image_subset_iff) | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2243 | from s show "s \<le> max 0 \<circ> f" | 
| 57137 | 2244 | by (subst s_eq) (auto simp: image_subset_iff le_fun_def split: split_indicator split_indicator_asm) | 
| 2245 | qed | |
| 2246 | then show ?thesis | |
| 2247 | unfolding nn_integral_def_finite SUP_def by simp | |
| 54417 | 2248 | qed | 
| 2249 | ||
| 59000 | 2250 | lemma nn_integral_count_space_indicator: | 
| 59779 | 2251 | assumes "NO_MATCH (UNIV::'a set) (X::'a set)" | 
| 59000 | 2252 | shows "(\<integral>\<^sup>+x. f x \<partial>count_space X) = (\<integral>\<^sup>+x. f x * indicator X x \<partial>count_space UNIV)" | 
| 2253 | by (simp add: nn_integral_restrict_space[symmetric] restrict_count_space) | |
| 2254 | ||
| 59425 | 2255 | lemma nn_integral_count_space_eq: | 
| 2256 | "(\<And>x. x \<in> A - B \<Longrightarrow> f x = 0) \<Longrightarrow> (\<And>x. x \<in> B - A \<Longrightarrow> f x = 0) \<Longrightarrow> | |
| 2257 | (\<integral>\<^sup>+x. f x \<partial>count_space A) = (\<integral>\<^sup>+x. f x \<partial>count_space B)" | |
| 2258 | by (auto simp: nn_integral_count_space_indicator intro!: nn_integral_cong split: split_indicator) | |
| 2259 | ||
| 59023 | 2260 | lemma nn_integral_ge_point: | 
| 2261 | assumes "x \<in> A" | |
| 2262 | shows "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space A" | |
| 2263 | proof - | |
| 2264 |   from assms have "p x \<le> \<integral>\<^sup>+ x. p x \<partial>count_space {x}"
 | |
| 2265 | by(auto simp add: nn_integral_count_space_finite max_def) | |
| 2266 |   also have "\<dots> = \<integral>\<^sup>+ x'. p x' * indicator {x} x' \<partial>count_space A"
 | |
| 2267 | using assms by(auto simp add: nn_integral_count_space_indicator indicator_def intro!: nn_integral_cong) | |
| 2268 | also have "\<dots> \<le> \<integral>\<^sup>+ x. max 0 (p x) \<partial>count_space A" | |
| 2269 | by(rule nn_integral_mono)(simp add: indicator_def) | |
| 2270 | also have "\<dots> = \<integral>\<^sup>+ x. p x \<partial>count_space A" by(simp add: nn_integral_def o_def) | |
| 2271 | finally show ?thesis . | |
| 2272 | qed | |
| 2273 | ||
| 61808 | 2274 | subsubsection \<open>Measure spaces with an associated density\<close> | 
| 47694 | 2275 | |
| 2276 | 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 | 2277 | "density M f = measure_of (space M) (sets M) (\<lambda>A. \<integral>\<^sup>+ x. f x * indicator A x \<partial>M)" | 
| 35582 | 2278 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2279 | lemma | 
| 59048 | 2280 | shows sets_density[simp, measurable_cong]: "sets (density M f) = sets M" | 
| 47694 | 2281 | and space_density[simp]: "space (density M f) = space M" | 
| 2282 | by (auto simp: density_def) | |
| 2283 | ||
| 50003 | 2284 | (* FIXME: add conversion to simplify space, sets and measurable *) | 
| 2285 | lemma space_density_imp[measurable_dest]: | |
| 2286 | "\<And>x M f. x \<in> space (density M f) \<Longrightarrow> x \<in> space M" by auto | |
| 2287 | ||
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2288 | lemma | 
| 47694 | 2289 | shows measurable_density_eq1[simp]: "g \<in> measurable (density Mg f) Mg' \<longleftrightarrow> g \<in> measurable Mg Mg'" | 
| 2290 | and measurable_density_eq2[simp]: "h \<in> measurable Mh (density Mh' f) \<longleftrightarrow> h \<in> measurable Mh Mh'" | |
| 2291 | and simple_function_density_eq[simp]: "simple_function (density Mu f) u \<longleftrightarrow> simple_function Mu u" | |
| 2292 | unfolding measurable_def simple_function_def by simp_all | |
| 2293 | ||
| 2294 | lemma density_cong: "f \<in> borel_measurable M \<Longrightarrow> f' \<in> borel_measurable M \<Longrightarrow> | |
| 2295 | (AE x in M. f x = f' x) \<Longrightarrow> density M f = density M f'" | |
| 56996 | 2296 | unfolding density_def by (auto intro!: measure_of_eq nn_integral_cong_AE sets.space_closed) | 
| 47694 | 2297 | |
| 2298 | lemma density_max_0: "density M f = density M (\<lambda>x. max 0 (f x))" | |
| 2299 | proof - | |
| 2300 | have "\<And>x A. max 0 (f x) * indicator A x = max 0 (f x * indicator A x)" | |
| 2301 | by (auto simp: indicator_def) | |
| 2302 | then show ?thesis | |
| 56996 | 2303 | unfolding density_def by (simp add: nn_integral_max_0) | 
| 47694 | 2304 | qed | 
| 2305 | ||
| 2306 | lemma density_ereal_max_0: "density M (\<lambda>x. ereal (f x)) = density M (\<lambda>x. ereal (max 0 (f x)))" | |
| 2307 | by (subst density_max_0) (auto intro!: arg_cong[where f="density M"] split: split_max) | |
| 38656 | 2308 | |
| 47694 | 2309 | lemma emeasure_density: | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 2310 | 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 | 2311 | shows "emeasure (density M f) A = (\<integral>\<^sup>+ x. f x * indicator A x \<partial>M)" | 
| 47694 | 2312 | (is "_ = ?\<mu> A") | 
| 2313 | unfolding density_def | |
| 2314 | proof (rule emeasure_measure_of_sigma) | |
| 2315 | show "sigma_algebra (space M) (sets M)" .. | |
| 2316 | show "positive (sets M) ?\<mu>" | |
| 56996 | 2317 | 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 | 2318 | have \<mu>_eq: "?\<mu> = (\<lambda>A. \<integral>\<^sup>+ x. max 0 (f x) * indicator A x \<partial>M)" (is "?\<mu> = ?\<mu>'") | 
| 56996 | 2319 | apply (subst nn_integral_max_0[symmetric]) | 
| 2320 | apply (intro ext nn_integral_cong_AE AE_I2) | |
| 47694 | 2321 | apply (auto simp: indicator_def) | 
| 2322 | done | |
| 2323 | show "countably_additive (sets M) ?\<mu>" | |
| 2324 | unfolding \<mu>_eq | |
| 2325 | proof (intro countably_additiveI) | |
| 2326 | 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 | 2327 | then have "\<And>i. A i \<in> sets M" by auto | 
| 47694 | 2328 | 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 | 2329 | by (auto simp: set_eq_iff) | 
| 47694 | 2330 | 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 | 2331 | have "(\<Sum>n. ?\<mu>' (A n)) = (\<integral>\<^sup>+ x. (\<Sum>n. max 0 (f x) * indicator (A n) x) \<partial>M)" | 
| 56996 | 2332 | 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 | 2333 | also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) * (\<Sum>n. indicator (A n) x) \<partial>M)" using f | 
| 56996 | 2334 | 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 | 2335 | also have "\<dots> = (\<integral>\<^sup>+ x. max 0 (f x) * indicator (\<Union>n. A n) x \<partial>M)" | 
| 47694 | 2336 | unfolding suminf_indicator[OF disj] .. | 
| 2337 | finally show "(\<Sum>n. ?\<mu>' (A n)) = ?\<mu>' (\<Union>x. A x)" by simp | |
| 2338 | qed | |
| 2339 | qed fact | |
| 38656 | 2340 | |
| 47694 | 2341 | lemma null_sets_density_iff: | 
| 2342 | assumes f: "f \<in> borel_measurable M" | |
| 2343 | 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)" | |
| 2344 | proof - | |
| 2345 |   { 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 | 2346 | 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 | 2347 | apply (subst nn_integral_max_0[symmetric]) | 
| 2348 | apply (intro nn_integral_cong) | |
| 47694 | 2349 | apply (auto simp: indicator_def) | 
| 2350 | done | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2351 | have "(\<integral>\<^sup>+x. f x * indicator A x \<partial>M) = 0 \<longleftrightarrow> | 
| 47694 | 2352 |       emeasure M {x \<in> space M. max 0 (f x) * indicator A x \<noteq> 0} = 0"
 | 
| 2353 | unfolding eq | |
| 61808 | 2354 | using f \<open>A \<in> sets M\<close> | 
| 56996 | 2355 | by (intro nn_integral_0_iff) auto | 
| 47694 | 2356 | also have "\<dots> \<longleftrightarrow> (AE x in M. max 0 (f x) * indicator A x = 0)" | 
| 61808 | 2357 | using f \<open>A \<in> sets M\<close> | 
| 50002 
ce0d316b5b44
add measurability prover; add support for Borel sets
 hoelzl parents: 
50001diff
changeset | 2358 | by (intro AE_iff_measurable[OF _ refl, symmetric]) auto | 
| 47694 | 2359 | 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)" | 
| 2360 | 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 | 2361 | 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 | 2362 | with f show ?thesis | 
| 2363 | by (simp add: null_sets_def emeasure_density cong: conj_cong) | |
| 2364 | qed | |
| 2365 | ||
| 2366 | lemma AE_density: | |
| 2367 | assumes f: "f \<in> borel_measurable M" | |
| 2368 | shows "(AE x in density M f. P x) \<longleftrightarrow> (AE x in M. 0 < f x \<longrightarrow> P x)" | |
| 2369 | proof | |
| 2370 | assume "AE x in density M f. P x" | |
| 2371 |   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"
 | |
| 2372 | by (auto simp: eventually_ae_filter null_sets_density_iff) | |
| 2373 | then have "AE x in M. x \<notin> N \<longrightarrow> P x" by auto | |
| 2374 | with ae show "AE x in M. 0 < f x \<longrightarrow> P x" | |
| 2375 | by (rule eventually_elim2) auto | |
| 2376 | next | |
| 2377 | fix N assume ae: "AE x in M. 0 < f x \<longrightarrow> P x" | |
| 2378 |   then obtain N where "{x \<in> space M. \<not> (0 < f x \<longrightarrow> P x)} \<subseteq> N" "N \<in> null_sets M"
 | |
| 2379 | by (auto simp: eventually_ae_filter) | |
| 2380 |   then have *: "{x \<in> space (density M f). \<not> P x} \<subseteq> N \<union> {x\<in>space M. \<not> 0 < f x}"
 | |
| 2381 |     "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 | 2382 | using f by (auto simp: subset_eq intro!: sets.sets_Collect_neg AE_not_in) | 
| 47694 | 2383 | show "AE x in density M f. P x" | 
| 2384 | using ae2 | |
| 2385 | unfolding eventually_ae_filter[of _ "density M f"] Bex_def null_sets_density_iff[OF f] | |
| 2386 |     by (intro exI[of _ "N \<union> {x\<in>space M. \<not> 0 < f x}"] conjI *)
 | |
| 2387 | (auto elim: eventually_elim2) | |
| 35582 | 2388 | qed | 
| 2389 | ||
| 56996 | 2390 | lemma nn_integral_density': | 
| 47694 | 2391 | 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 | 2392 | assumes g: "g \<in> borel_measurable M" "\<And>x. 0 \<le> g x" | 
| 56996 | 2393 | shows "integral\<^sup>N (density M f) g = (\<integral>\<^sup>+ x. f x * g x \<partial>M)" | 
| 49798 | 2394 | using g proof induct | 
| 2395 | case (cong u v) | |
| 49799 
15ea98537c76
strong nonnegativ (instead of ae nn) for induction rule
 hoelzl parents: 
49798diff
changeset | 2396 | then show ?case | 
| 56996 | 2397 | apply (subst nn_integral_cong[OF cong(3)]) | 
| 2398 | apply (simp_all cong: nn_integral_cong) | |
| 49798 | 2399 | done | 
| 2400 | next | |
| 2401 | case (set A) then show ?case | |
| 2402 | by (simp add: emeasure_density f) | |
| 2403 | next | |
| 2404 | case (mult u c) | |
| 2405 | moreover have "\<And>x. f x * (c * u x) = c * (f x * u x)" by (simp add: field_simps) | |
| 2406 | ultimately show ?case | |
| 56996 | 2407 | using f by (simp add: nn_integral_cmult) | 
| 49798 | 2408 | next | 
| 2409 | case (add u v) | |
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 2410 | then have "\<And>x. f x * (v x + u x) = f x * v x + f x * u x" | 
| 49798 | 2411 | by (simp add: ereal_right_distrib) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 2412 | with add f show ?case | 
| 56996 | 2413 | by (auto simp add: nn_integral_add ereal_zero_le_0_iff intro!: nn_integral_add[symmetric]) | 
| 49798 | 2414 | next | 
| 2415 | case (seq U) | |
| 2416 | from f(2) have eq: "AE x in M. f x * (SUP i. U i x) = (SUP i. f x * U i x)" | |
| 59452 
2538b2c51769
ereal: tuned proofs concerning continuity and suprema
 hoelzl parents: 
59426diff
changeset | 2417 | by eventually_elim (simp add: SUP_ereal_mult_left seq) | 
| 49798 | 2418 | from seq f show ?case | 
| 56996 | 2419 | apply (simp add: nn_integral_monotone_convergence_SUP) | 
| 2420 | apply (subst nn_integral_cong_AE[OF eq]) | |
| 2421 | apply (subst nn_integral_monotone_convergence_SUP_AE) | |
| 49798 | 2422 | apply (auto simp: incseq_def le_fun_def intro!: ereal_mult_left_mono) | 
| 2423 | done | |
| 47694 | 2424 | qed | 
| 38705 | 2425 | |
| 56996 | 2426 | lemma nn_integral_density: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2427 | "f \<in> borel_measurable M \<Longrightarrow> AE x in M. 0 \<le> f x \<Longrightarrow> g' \<in> borel_measurable M \<Longrightarrow> | 
| 56996 | 2428 | integral\<^sup>N (density M f) g' = (\<integral>\<^sup>+ x. f x * g' x \<partial>M)" | 
| 2429 | by (subst (1 2) nn_integral_max_0[symmetric]) | |
| 2430 | (auto intro!: nn_integral_cong_AE | |
| 2431 | simp: measurable_If max_def ereal_zero_le_0_iff nn_integral_density') | |
| 49798 | 2432 | |
| 57275 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2433 | lemma density_distr: | 
| 
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
 hoelzl parents: 
57137diff
changeset | 2434 | 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 | 2435 | 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 | 2436 | 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 | 2437 | (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 | 2438 | 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 | 2439 | |
| 47694 | 2440 | lemma emeasure_restricted: | 
| 2441 | assumes S: "S \<in> sets M" and X: "X \<in> sets M" | |
| 2442 | shows "emeasure (density M (indicator S)) X = emeasure M (S \<inter> X)" | |
| 38705 | 2443 | proof - | 
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2444 | have "emeasure (density M (indicator S)) X = (\<integral>\<^sup>+x. indicator S x * indicator X x \<partial>M)" | 
| 47694 | 2445 | 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 | 2446 | also have "\<dots> = (\<integral>\<^sup>+x. indicator (S \<inter> X) x \<partial>M)" | 
| 56996 | 2447 | by (auto intro!: nn_integral_cong simp: indicator_def) | 
| 47694 | 2448 | also have "\<dots> = emeasure M (S \<inter> X)" | 
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 2449 | using S X by (simp add: sets.Int) | 
| 47694 | 2450 | finally show ?thesis . | 
| 2451 | qed | |
| 2452 | ||
| 2453 | lemma measure_restricted: | |
| 2454 | "S \<in> sets M \<Longrightarrow> X \<in> sets M \<Longrightarrow> measure (density M (indicator S)) X = measure M (S \<inter> X)" | |
| 2455 | by (simp add: emeasure_restricted measure_def) | |
| 2456 | ||
| 2457 | lemma (in finite_measure) finite_measure_restricted: | |
| 2458 | "S \<in> sets M \<Longrightarrow> finite_measure (density M (indicator S))" | |
| 61169 | 2459 | by standard (simp add: emeasure_restricted) | 
| 47694 | 2460 | |
| 2461 | lemma emeasure_density_const: | |
| 2462 | "A \<in> sets M \<Longrightarrow> 0 \<le> c \<Longrightarrow> emeasure (density M (\<lambda>_. c)) A = c * emeasure M A" | |
| 56996 | 2463 | by (auto simp: nn_integral_cmult_indicator emeasure_density) | 
| 47694 | 2464 | |
| 2465 | lemma measure_density_const: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2466 | "A \<in> sets M \<Longrightarrow> 0 \<le> c \<Longrightarrow> c \<noteq> \<infinity> \<Longrightarrow> measure (density M (\<lambda>_. c)) A = real_of_ereal c * measure M A" | 
| 47694 | 2467 | by (auto simp: emeasure_density_const measure_def) | 
| 2468 | ||
| 2469 | lemma density_density_eq: | |
| 2470 | "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow> AE x in M. 0 \<le> f x \<Longrightarrow> | |
| 2471 | density (density M f) g = density M (\<lambda>x. f x * g x)" | |
| 56996 | 2472 | by (auto intro!: measure_eqI simp: emeasure_density nn_integral_density ac_simps) | 
| 47694 | 2473 | |
| 2474 | lemma distr_density_distr: | |
| 2475 | assumes T: "T \<in> measurable M M'" and T': "T' \<in> measurable M' M" | |
| 2476 | and inv: "\<forall>x\<in>space M. T' (T x) = x" | |
| 2477 | assumes f: "f \<in> borel_measurable M'" | |
| 2478 | shows "distr (density (distr M M' T) f) M T' = density M (f \<circ> T)" (is "?R = ?L") | |
| 2479 | proof (rule measure_eqI) | |
| 2480 | fix A assume A: "A \<in> sets ?R" | |
| 2481 |   { fix x assume "x \<in> space M"
 | |
| 50244 
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
 immler parents: 
50104diff
changeset | 2482 | with sets.sets_into_space[OF A] | 
| 47694 | 2483 | have "indicator (T' -` A \<inter> space M') (T x) = (indicator A x :: ereal)" | 
| 2484 | using T inv by (auto simp: indicator_def measurable_space) } | |
| 2485 | with A T T' f show "emeasure ?R A = emeasure ?L A" | |
| 2486 | by (simp add: measurable_comp emeasure_density emeasure_distr | |
| 56996 | 2487 | nn_integral_distr measurable_sets cong: nn_integral_cong) | 
| 47694 | 2488 | qed simp | 
| 2489 | ||
| 2490 | lemma density_density_divide: | |
| 2491 | fixes f g :: "'a \<Rightarrow> real" | |
| 2492 | assumes f: "f \<in> borel_measurable M" "AE x in M. 0 \<le> f x" | |
| 2493 | assumes g: "g \<in> borel_measurable M" "AE x in M. 0 \<le> g x" | |
| 2494 | assumes ac: "AE x in M. f x = 0 \<longrightarrow> g x = 0" | |
| 2495 | shows "density (density M f) (\<lambda>x. g x / f x) = density M g" | |
| 2496 | proof - | |
| 2497 | have "density M g = density M (\<lambda>x. f x * (g x / f x))" | |
| 2498 | using f g ac by (auto intro!: density_cong measurable_If) | |
| 2499 | then show ?thesis | |
| 2500 | using f g by (subst density_density_eq) auto | |
| 38705 | 2501 | qed | 
| 2502 | ||
| 59425 | 2503 | lemma density_1: "density M (\<lambda>_. 1) = M" | 
| 2504 | by (intro measure_eqI) (auto simp: emeasure_density) | |
| 2505 | ||
| 2506 | lemma emeasure_density_add: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2507 | assumes X: "X \<in> sets M" | 
| 59425 | 2508 | assumes Mf[measurable]: "f \<in> borel_measurable M" | 
| 2509 | assumes Mg[measurable]: "g \<in> borel_measurable M" | |
| 2510 | assumes nonnegf: "\<And>x. x \<in> space M \<Longrightarrow> f x \<ge> 0" | |
| 2511 | assumes nonnegg: "\<And>x. x \<in> space M \<Longrightarrow> g x \<ge> 0" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2512 | shows "emeasure (density M f) X + emeasure (density M g) X = | 
| 59425 | 2513 | emeasure (density M (\<lambda>x. f x + g x)) X" | 
| 2514 | using assms | |
| 2515 | apply (subst (1 2 3) emeasure_density, simp_all) [] | |
| 2516 | apply (subst nn_integral_add[symmetric], simp_all) [] | |
| 2517 | apply (intro nn_integral_cong, simp split: split_indicator) | |
| 2518 | done | |
| 2519 | ||
| 61808 | 2520 | subsubsection \<open>Point measure\<close> | 
| 47694 | 2521 | |
| 2522 | definition point_measure :: "'a set \<Rightarrow> ('a \<Rightarrow> ereal) \<Rightarrow> 'a measure" where
 | |
| 2523 | "point_measure A f = density (count_space A) f" | |
| 2524 | ||
| 2525 | lemma | |
| 2526 | shows space_point_measure: "space (point_measure A f) = A" | |
| 2527 | and sets_point_measure: "sets (point_measure A f) = Pow A" | |
| 2528 | by (auto simp: point_measure_def) | |
| 2529 | ||
| 59048 | 2530 | lemma sets_point_measure_count_space[measurable_cong]: "sets (point_measure A f) = sets (count_space A)" | 
| 2531 | by (simp add: sets_point_measure) | |
| 2532 | ||
| 47694 | 2533 | lemma measurable_point_measure_eq1[simp]: | 
| 2534 | "g \<in> measurable (point_measure A f) M \<longleftrightarrow> g \<in> A \<rightarrow> space M" | |
| 2535 | unfolding point_measure_def by simp | |
| 2536 | ||
| 2537 | lemma measurable_point_measure_eq2_finite[simp]: | |
| 2538 | "finite A \<Longrightarrow> | |
| 2539 | g \<in> measurable M (point_measure A f) \<longleftrightarrow> | |
| 2540 |     (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 | 2541 | unfolding point_measure_def by (simp add: measurable_count_space_eq2) | 
| 47694 | 2542 | |
| 2543 | lemma simple_function_point_measure[simp]: | |
| 2544 | "simple_function (point_measure A f) g \<longleftrightarrow> finite (g ` A)" | |
| 2545 | by (simp add: point_measure_def) | |
| 2546 | ||
| 2547 | lemma emeasure_point_measure: | |
| 2548 |   assumes A: "finite {a\<in>X. 0 < f a}" "X \<subseteq> A"
 | |
| 2549 | shows "emeasure (point_measure A f) X = (\<Sum>a|a\<in>X \<and> 0 < f a. f a)" | |
| 35977 | 2550 | proof - | 
| 47694 | 2551 |   have "{a. (a \<in> X \<longrightarrow> a \<in> A \<and> 0 < f a) \<and> a \<in> X} = {a\<in>X. 0 < f a}"
 | 
| 61808 | 2552 | using \<open>X \<subseteq> A\<close> by auto | 
| 47694 | 2553 | with A show ?thesis | 
| 56996 | 2554 | by (simp add: emeasure_density nn_integral_count_space ereal_zero_le_0_iff | 
| 47694 | 2555 | point_measure_def indicator_def) | 
| 35977 | 2556 | qed | 
| 2557 | ||
| 47694 | 2558 | lemma emeasure_point_measure_finite: | 
| 49795 | 2559 | "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 | 2560 | by (subst emeasure_point_measure) (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le) | 
| 47694 | 2561 | |
| 49795 | 2562 | lemma emeasure_point_measure_finite2: | 
| 2563 | "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)" | |
| 2564 | by (subst emeasure_point_measure) | |
| 57418 | 2565 | (auto dest: finite_subset intro!: setsum.mono_neutral_left simp: less_le) | 
| 49795 | 2566 | |
| 47694 | 2567 | lemma null_sets_point_measure_iff: | 
| 2568 | "X \<in> null_sets (point_measure A f) \<longleftrightarrow> X \<subseteq> A \<and> (\<forall>x\<in>X. f x \<le> 0)" | |
| 2569 | by (auto simp: AE_count_space null_sets_density_iff point_measure_def) | |
| 2570 | ||
| 2571 | lemma AE_point_measure: | |
| 2572 | "(AE x in point_measure A f. P x) \<longleftrightarrow> (\<forall>x\<in>A. 0 < f x \<longrightarrow> P x)" | |
| 2573 | unfolding point_measure_def | |
| 2574 | by (subst AE_density) (auto simp: AE_density AE_count_space point_measure_def) | |
| 2575 | ||
| 56996 | 2576 | lemma nn_integral_point_measure: | 
| 47694 | 2577 |   "finite {a\<in>A. 0 < f a \<and> 0 < g a} \<Longrightarrow>
 | 
| 56996 | 2578 | 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 | 2579 | unfolding point_measure_def | 
| 2580 | apply (subst density_max_0) | |
| 56996 | 2581 | apply (subst nn_integral_density) | 
| 2582 | apply (simp_all add: AE_count_space nn_integral_density) | |
| 2583 | apply (subst nn_integral_count_space ) | |
| 57418 | 2584 | apply (auto intro!: setsum.cong simp: max_def ereal_zero_less_0_iff) | 
| 47694 | 2585 | apply (rule finite_subset) | 
| 2586 | prefer 2 | |
| 2587 | apply assumption | |
| 2588 | apply auto | |
| 2589 | done | |
| 2590 | ||
| 56996 | 2591 | lemma nn_integral_point_measure_finite: | 
| 47694 | 2592 | "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 | 2593 | integral\<^sup>N (point_measure A f) g = (\<Sum>a\<in>A. f a * g a)" | 
| 57418 | 2594 | by (subst nn_integral_point_measure) (auto intro!: setsum.mono_neutral_left simp: less_le) | 
| 47694 | 2595 | |
| 61808 | 2596 | subsubsection \<open>Uniform measure\<close> | 
| 47694 | 2597 | |
| 2598 | definition "uniform_measure M A = density M (\<lambda>x. indicator A x / emeasure M A)" | |
| 2599 | ||
| 2600 | lemma | |
| 59048 | 2601 | shows sets_uniform_measure[simp, measurable_cong]: "sets (uniform_measure M A) = sets M" | 
| 47694 | 2602 | and space_uniform_measure[simp]: "space (uniform_measure M A) = space M" | 
| 2603 | by (auto simp: uniform_measure_def) | |
| 2604 | ||
| 2605 | lemma emeasure_uniform_measure[simp]: | |
| 2606 | assumes A: "A \<in> sets M" and B: "B \<in> sets M" | |
| 2607 | shows "emeasure (uniform_measure M A) B = emeasure M (A \<inter> B) / emeasure M A" | |
| 2608 | proof - | |
| 53015 
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
 wenzelm parents: 
51340diff
changeset | 2609 | 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 | 2610 | by (auto simp add: uniform_measure_def emeasure_density split: split_indicator | 
| 56996 | 2611 | intro!: nn_integral_cong) | 
| 47694 | 2612 | also have "\<dots> = emeasure M (A \<inter> B) / emeasure M A" | 
| 2613 | using A B | |
| 56996 | 2614 | by (subst nn_integral_cmult_indicator) (simp_all add: sets.Int emeasure_nonneg) | 
| 47694 | 2615 | finally show ?thesis . | 
| 2616 | qed | |
| 2617 | ||
| 2618 | lemma measure_uniform_measure[simp]: | |
| 2619 | assumes A: "emeasure M A \<noteq> 0" "emeasure M A \<noteq> \<infinity>" and B: "B \<in> sets M" | |
| 2620 | shows "measure (uniform_measure M A) B = measure M (A \<inter> B) / measure M A" | |
| 2621 | using emeasure_uniform_measure[OF emeasure_neq_0_sets[OF A(1)] B] A | |
| 2622 | by (cases "emeasure M A" "emeasure M (A \<inter> B)" rule: ereal2_cases) (simp_all add: measure_def) | |
| 2623 | ||
| 58606 | 2624 | lemma AE_uniform_measureI: | 
| 2625 | "A \<in> sets M \<Longrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x) \<Longrightarrow> (AE x in uniform_measure M A. P x)" | |
| 2626 | unfolding uniform_measure_def by (auto simp: AE_density) | |
| 2627 | ||
| 59000 | 2628 | lemma emeasure_uniform_measure_1: | 
| 2629 | "emeasure M S \<noteq> 0 \<Longrightarrow> emeasure M S \<noteq> \<infinity> \<Longrightarrow> emeasure (uniform_measure M S) S = 1" | |
| 2630 | by (subst emeasure_uniform_measure) | |
| 2631 | (simp_all add: emeasure_nonneg emeasure_neq_0_sets) | |
| 2632 | ||
| 2633 | lemma nn_integral_uniform_measure: | |
| 2634 | assumes f[measurable]: "f \<in> borel_measurable M" and "\<And>x. 0 \<le> f x" and S[measurable]: "S \<in> sets M" | |
| 2635 | shows "(\<integral>\<^sup>+x. f x \<partial>uniform_measure M S) = (\<integral>\<^sup>+x. f x * indicator S x \<partial>M) / emeasure M S" | |
| 2636 | proof - | |
| 2637 |   { assume "emeasure M S = \<infinity>"
 | |
| 2638 | then have ?thesis | |
| 2639 | by (simp add: uniform_measure_def nn_integral_density f) } | |
| 2640 | moreover | |
| 2641 |   { assume [simp]: "emeasure M S = 0"
 | |
| 2642 | then have ae: "AE x in M. x \<notin> S" | |
| 2643 | using sets.sets_into_space[OF S] | |
| 2644 | by (subst AE_iff_measurable[OF _ refl]) (simp_all add: subset_eq cong: rev_conj_cong) | |
| 2645 | from ae have "(\<integral>\<^sup>+ x. indicator S x * f x / 0 \<partial>M) = 0" | |
| 2646 | by (subst nn_integral_0_iff_AE) auto | |
| 2647 | moreover from ae have "(\<integral>\<^sup>+ x. f x * indicator S x \<partial>M) = 0" | |
| 2648 | by (subst nn_integral_0_iff_AE) auto | |
| 2649 | ultimately have ?thesis | |
| 2650 | by (simp add: uniform_measure_def nn_integral_density f) } | |
| 2651 | moreover | |
| 2652 |   { assume "emeasure M S \<noteq> 0" "emeasure M S \<noteq> \<infinity>"
 | |
| 2653 | moreover then have "0 < emeasure M S" | |
| 2654 | by (simp add: emeasure_nonneg less_le) | |
| 2655 | ultimately have ?thesis | |
| 2656 | unfolding uniform_measure_def | |
| 2657 | apply (subst nn_integral_density) | |
| 2658 | apply (auto simp: f nn_integral_divide intro!: zero_le_divide_ereal) | |
| 2659 | apply (simp add: mult.commute) | |
| 2660 | done } | |
| 2661 | ultimately show ?thesis by blast | |
| 2662 | qed | |
| 2663 | ||
| 2664 | lemma AE_uniform_measure: | |
| 2665 | assumes "emeasure M A \<noteq> 0" "emeasure M A < \<infinity>" | |
| 2666 | shows "(AE x in uniform_measure M A. P x) \<longleftrightarrow> (AE x in M. x \<in> A \<longrightarrow> P x)" | |
| 2667 | proof - | |
| 2668 | have "A \<in> sets M" | |
| 61808 | 2669 | using \<open>emeasure M A \<noteq> 0\<close> by (metis emeasure_notin_sets) | 
| 59000 | 2670 | moreover have "\<And>x. 0 < indicator A x / emeasure M A \<longleftrightarrow> x \<in> A" | 
| 2671 | using emeasure_nonneg[of M A] assms | |
| 2672 | by (cases "emeasure M A") (auto split: split_indicator simp: one_ereal_def) | |
| 2673 | ultimately show ?thesis | |
| 2674 | unfolding uniform_measure_def by (simp add: AE_density) | |
| 2675 | qed | |
| 2676 | ||
| 61808 | 2677 | subsubsection \<open>Null measure\<close> | 
| 59425 | 2678 | |
| 2679 | lemma null_measure_eq_density: "null_measure M = density M (\<lambda>_. 0)" | |
| 2680 | by (intro measure_eqI) (simp_all add: emeasure_density) | |
| 2681 | ||
| 2682 | lemma nn_integral_null_measure[simp]: "(\<integral>\<^sup>+x. f x \<partial>null_measure M) = 0" | |
| 2683 | by (auto simp add: nn_integral_def simple_integral_def SUP_constant bot_ereal_def le_fun_def | |
| 2684 | intro!: exI[of _ "\<lambda>x. 0"]) | |
| 2685 | ||
| 2686 | lemma density_null_measure[simp]: "density (null_measure M) f = null_measure M" | |
| 2687 | proof (intro measure_eqI) | |
| 2688 | fix A show "emeasure (density (null_measure M) f) A = emeasure (null_measure M) A" | |
| 2689 | by (simp add: density_def) (simp only: null_measure_def[symmetric] emeasure_null_measure) | |
| 2690 | qed simp | |
| 2691 | ||
| 61808 | 2692 | subsubsection \<open>Uniform count measure\<close> | 
| 47694 | 2693 | |
| 2694 | definition "uniform_count_measure A = point_measure A (\<lambda>x. 1 / card A)" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2695 | |
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2696 | lemma | 
| 47694 | 2697 | shows space_uniform_count_measure: "space (uniform_count_measure A) = A" | 
| 2698 | and sets_uniform_count_measure: "sets (uniform_count_measure A) = Pow A" | |
| 2699 | unfolding uniform_count_measure_def by (auto simp: space_point_measure sets_point_measure) | |
| 59048 | 2700 | |
| 2701 | lemma sets_uniform_count_measure_count_space[measurable_cong]: | |
| 2702 | "sets (uniform_count_measure A) = sets (count_space A)" | |
| 2703 | by (simp add: sets_uniform_count_measure) | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2704 | |
| 47694 | 2705 | lemma emeasure_uniform_count_measure: | 
| 2706 | "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> emeasure (uniform_count_measure A) X = card X / card A" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2707 | by (simp add: emeasure_point_measure_finite uniform_count_measure_def) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2708 | |
| 47694 | 2709 | lemma measure_uniform_count_measure: | 
| 2710 | "finite A \<Longrightarrow> X \<subseteq> A \<Longrightarrow> measure (uniform_count_measure A) X = card X / card A" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61359diff
changeset | 2711 | by (simp add: emeasure_point_measure_finite uniform_count_measure_def measure_def) | 
| 47694 | 2712 | |
| 61633 | 2713 | lemma space_uniform_count_measure_empty_iff [simp]: | 
| 2714 |   "space (uniform_count_measure X) = {} \<longleftrightarrow> X = {}"
 | |
| 2715 | by(simp add: space_uniform_count_measure) | |
| 2716 | ||
| 2717 | lemma sets_uniform_count_measure_eq_UNIV [simp]: | |
| 2718 | "sets (uniform_count_measure UNIV) = UNIV \<longleftrightarrow> True" | |
| 2719 | "UNIV = sets (uniform_count_measure UNIV) \<longleftrightarrow> True" | |
| 2720 | by(simp_all add: sets_uniform_count_measure) | |
| 2721 | ||
| 61634 | 2722 | subsubsection \<open>Scaled measure\<close> | 
| 2723 | ||
| 2724 | lemma nn_integral_scale_measure': | |
| 2725 | assumes f: "f \<in> borel_measurable M" "\<And>x. 0 \<le> f x" | |
| 2726 | shows "nn_integral (scale_measure r M) f = max 0 r * nn_integral M f" | |
| 2727 | using f | |
| 2728 | proof induction | |
| 2729 | case (cong f g) | |
| 2730 | thus ?case | |
| 2731 | by(simp add: cong.hyps space_scale_measure cong: nn_integral_cong_simp) | |
| 2732 | next | |
| 2733 | case (mult f c) | |
| 2734 | thus ?case | |
| 2735 | by(simp add: nn_integral_cmult max_def mult.assoc mult.left_commute) | |
| 2736 | next | |
| 2737 | case (add f g) | |
| 2738 | thus ?case | |
| 2739 | by(simp add: nn_integral_add ereal_right_distrib nn_integral_nonneg) | |
| 2740 | next | |
| 2741 | case (seq U) | |
| 2742 | thus ?case | |
| 2743 | by(simp add: nn_integral_monotone_convergence_SUP SUP_ereal_mult_left nn_integral_nonneg) | |
| 2744 | qed simp | |
| 2745 | ||
| 2746 | lemma nn_integral_scale_measure: | |
| 2747 | "f \<in> borel_measurable M \<Longrightarrow> nn_integral (scale_measure r M) f = max 0 r * nn_integral M f" | |
| 2748 | by(subst (1 2) nn_integral_max_0[symmetric])(rule nn_integral_scale_measure', simp_all) | |
| 2749 | ||
| 35748 | 2750 | end |