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