| author | wenzelm | 
| Thu, 02 Aug 2012 13:37:58 +0200 | |
| changeset 48647 | a5144c4c26a2 | 
| parent 47757 | 5e6fe71e2390 | 
| child 49777 | 6ac97ab9b295 | 
| permissions | -rw-r--r-- | 
| 42067 | 1 | (* Title: HOL/Probability/Lebesgue_Measure.thy | 
| 2 | Author: Johannes Hölzl, TU München | |
| 3 | Author: Robert Himmelmann, TU München | |
| 4 | *) | |
| 5 | ||
| 38656 | 6 | header {* Lebsegue measure *}
 | 
| 42067 | 7 | |
| 38656 | 8 | theory Lebesgue_Measure | 
| 42146 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 hoelzl parents: 
42067diff
changeset | 9 | imports Finite_Product_Measure | 
| 38656 | 10 | begin | 
| 11 | ||
| 47694 | 12 | lemma borel_measurable_sets: | 
| 13 | assumes "f \<in> measurable borel M" "A \<in> sets M" | |
| 14 | shows "f -` A \<in> sets borel" | |
| 15 | using measurable_sets[OF assms] by simp | |
| 16 | ||
| 17 | lemma measurable_identity[intro,simp]: | |
| 18 | "(\<lambda>x. x) \<in> measurable M M" | |
| 19 | unfolding measurable_def by auto | |
| 20 | ||
| 38656 | 21 | subsection {* Standard Cubes *}
 | 
| 22 | ||
| 40859 | 23 | definition cube :: "nat \<Rightarrow> 'a::ordered_euclidean_space set" where | 
| 24 |   "cube n \<equiv> {\<chi>\<chi> i. - real n .. \<chi>\<chi> i. real n}"
 | |
| 25 | ||
| 26 | lemma cube_closed[intro]: "closed (cube n)" | |
| 27 | unfolding cube_def by auto | |
| 28 | ||
| 29 | lemma cube_subset[intro]: "n \<le> N \<Longrightarrow> cube n \<subseteq> cube N" | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44666diff
changeset | 30 | by (fastforce simp: eucl_le[where 'a='a] cube_def) | 
| 38656 | 31 | |
| 40859 | 32 | lemma cube_subset_iff: | 
| 33 | "cube n \<subseteq> cube N \<longleftrightarrow> n \<le> N" | |
| 34 | proof | |
| 35 | assume subset: "cube n \<subseteq> (cube N::'a set)" | |
| 36 | then have "((\<chi>\<chi> i. real n)::'a) \<in> cube N" | |
| 37 | using DIM_positive[where 'a='a] | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44666diff
changeset | 38 | by (fastforce simp: cube_def eucl_le[where 'a='a]) | 
| 40859 | 39 | then show "n \<le> N" | 
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44666diff
changeset | 40 | by (fastforce simp: cube_def eucl_le[where 'a='a]) | 
| 40859 | 41 | next | 
| 42 | assume "n \<le> N" then show "cube n \<subseteq> (cube N::'a set)" by (rule cube_subset) | |
| 43 | qed | |
| 38656 | 44 | |
| 45 | lemma ball_subset_cube:"ball (0::'a::ordered_euclidean_space) (real n) \<subseteq> cube n" | |
| 46 | unfolding cube_def subset_eq mem_interval apply safe unfolding euclidean_lambda_beta' | |
| 47 | proof- fix x::'a and i assume as:"x\<in>ball 0 (real n)" "i<DIM('a)"
 | |
| 48 | thus "- real n \<le> x $$ i" "real n \<ge> x $$ i" | |
| 49 | using component_le_norm[of x i] by(auto simp: dist_norm) | |
| 50 | qed | |
| 51 | ||
| 52 | lemma mem_big_cube: obtains n where "x \<in> cube n" | |
| 44666 | 53 | proof- from reals_Archimedean2[of "norm x"] guess n .. | 
| 38656 | 54 | thus ?thesis apply-apply(rule that[where n=n]) | 
| 55 | apply(rule ball_subset_cube[unfolded subset_eq,rule_format]) | |
| 56 | by (auto simp add:dist_norm) | |
| 57 | qed | |
| 58 | ||
| 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 | 59 | lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)" | 
| 46905 | 60 | unfolding cube_def subset_eq apply safe unfolding mem_interval apply auto done | 
| 41654 | 61 | |
| 47757 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 62 | subsection {* Lebesgue measure *}
 | 
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 63 | |
| 47694 | 64 | definition lebesgue :: "'a::ordered_euclidean_space measure" where | 
| 65 |   "lebesgue = measure_of UNIV {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n}
 | |
| 66 | (\<lambda>A. SUP n. ereal (integral (cube n) (indicator A)))" | |
| 41661 | 67 | |
| 41654 | 68 | lemma space_lebesgue[simp]: "space lebesgue = UNIV" | 
| 69 | unfolding lebesgue_def by simp | |
| 70 | ||
| 71 | lemma lebesgueI: "(\<And>n. (indicator A :: _ \<Rightarrow> real) integrable_on cube n) \<Longrightarrow> A \<in> sets lebesgue" | |
| 72 | unfolding lebesgue_def by simp | |
| 73 | ||
| 74 | lemma absolutely_integrable_on_indicator[simp]: | |
| 75 | fixes A :: "'a::ordered_euclidean_space set" | |
| 76 | shows "((indicator A :: _ \<Rightarrow> real) absolutely_integrable_on X) \<longleftrightarrow> | |
| 77 | (indicator A :: _ \<Rightarrow> real) integrable_on X" | |
| 78 | unfolding absolutely_integrable_on_def by simp | |
| 79 | ||
| 80 | lemma LIMSEQ_indicator_UN: | |
| 81 | "(\<lambda>k. indicator (\<Union> i<k. A i) x) ----> (indicator (\<Union>i. A i) x :: real)" | |
| 82 | proof cases | |
| 83 | assume "\<exists>i. x \<in> A i" then guess i .. note i = this | |
| 84 | then have *: "\<And>n. (indicator (\<Union> i<n + Suc i. A i) x :: real) = 1" | |
| 85 | "(indicator (\<Union> i. A i) x :: real) = 1" by (auto simp: indicator_def) | |
| 86 | show ?thesis | |
| 87 | apply (rule LIMSEQ_offset[of _ "Suc i"]) unfolding * by auto | |
| 88 | qed (auto simp: indicator_def) | |
| 38656 | 89 | |
| 41654 | 90 | lemma indicator_add: | 
| 91 |   "A \<inter> B = {} \<Longrightarrow> (indicator A x::_::monoid_add) + indicator B x = indicator (A \<union> B) x"
 | |
| 92 | unfolding indicator_def by auto | |
| 38656 | 93 | |
| 47694 | 94 | lemma sigma_algebra_lebesgue: | 
| 95 |   defines "leb \<equiv> {A. \<forall>n. (indicator A :: 'a::ordered_euclidean_space \<Rightarrow> real) integrable_on cube n}"
 | |
| 96 | shows "sigma_algebra UNIV leb" | |
| 97 | proof (safe intro!: sigma_algebra_iff2[THEN iffD2]) | |
| 98 | fix A assume A: "A \<in> leb" | |
| 99 | moreover have "indicator (UNIV - A) = (\<lambda>x. 1 - indicator A x :: real)" | |
| 41654 | 100 | by (auto simp: fun_eq_iff indicator_def) | 
| 47694 | 101 | ultimately show "UNIV - A \<in> leb" | 
| 102 | using A by (auto intro!: integrable_sub simp: cube_def leb_def) | |
| 41654 | 103 | next | 
| 47694 | 104 |   fix n show "{} \<in> leb"
 | 
| 105 | by (auto simp: cube_def indicator_def[abs_def] leb_def) | |
| 41654 | 106 | next | 
| 47694 | 107 | fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> leb" | 
| 108 | have "\<forall>n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n" (is "\<forall>n. ?g integrable_on _") | |
| 109 | proof (intro dominated_convergence[where g="?g"] ballI allI) | |
| 110 | fix k n show "(indicator (\<Union>i<k. A i) :: _ \<Rightarrow> real) integrable_on cube n" | |
| 41654 | 111 | proof (induct k) | 
| 112 | case (Suc k) | |
| 113 | have *: "(\<Union> i<Suc k. A i) = (\<Union> i<k. A i) \<union> A k" | |
| 114 | unfolding lessThan_Suc UN_insert by auto | |
| 115 | have *: "(\<lambda>x. max (indicator (\<Union> i<k. A i) x) (indicator (A k) x) :: real) = | |
| 116 | indicator (\<Union> i<Suc k. A i)" (is "(\<lambda>x. max (?f x) (?g x)) = _") | |
| 117 | by (auto simp: fun_eq_iff * indicator_def) | |
| 118 | show ?case | |
| 47694 | 119 | using absolutely_integrable_max[of ?f "cube n" ?g] A Suc | 
| 120 | by (simp add: * leb_def subset_eq) | |
| 41654 | 121 | qed auto | 
| 122 | qed (auto intro: LIMSEQ_indicator_UN simp: cube_def) | |
| 47694 | 123 | then show "(\<Union>i. A i) \<in> leb" by (auto simp: leb_def) | 
| 41654 | 124 | qed simp | 
| 38656 | 125 | |
| 47694 | 126 | lemma sets_lebesgue: "sets lebesgue = {A. \<forall>n. (indicator A :: _ \<Rightarrow> real) integrable_on cube n}"
 | 
| 127 | unfolding lebesgue_def sigma_algebra.sets_measure_of_eq[OF sigma_algebra_lebesgue] .. | |
| 128 | ||
| 129 | lemma lebesgueD: "A \<in> sets lebesgue \<Longrightarrow> (indicator A :: _ \<Rightarrow> real) integrable_on cube n" | |
| 130 | unfolding sets_lebesgue by simp | |
| 131 | ||
| 47757 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 132 | lemma emeasure_lebesgue: | 
| 47694 | 133 | assumes "A \<in> sets lebesgue" | 
| 134 | shows "emeasure lebesgue A = (SUP n. ereal (integral (cube n) (indicator A)))" | |
| 135 | (is "_ = ?\<mu> A") | |
| 136 | proof (rule emeasure_measure_of[OF lebesgue_def]) | |
| 41654 | 137 |   have *: "indicator {} = (\<lambda>x. 0 :: real)" by (simp add: fun_eq_iff)
 | 
| 47694 | 138 | show "positive (sets lebesgue) ?\<mu>" | 
| 139 | proof (unfold positive_def, intro conjI ballI) | |
| 140 |     show "?\<mu> {} = 0" by (simp add: integral_0 *)
 | |
| 141 | fix A :: "'a set" assume "A \<in> sets lebesgue" then show "0 \<le> ?\<mu> A" | |
| 142 | by (auto intro!: SUP_upper2 Integration.integral_nonneg simp: sets_lebesgue) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 143 | qed | 
| 40859 | 144 | next | 
| 47694 | 145 | show "countably_additive (sets lebesgue) ?\<mu>" | 
| 41654 | 146 | proof (intro countably_additive_def[THEN iffD2] allI impI) | 
| 47694 | 147 | fix A :: "nat \<Rightarrow> 'a set" assume rA: "range A \<subseteq> sets lebesgue" "disjoint_family A" | 
| 41654 | 148 | then have A[simp, intro]: "\<And>i n. (indicator (A i) :: _ \<Rightarrow> real) integrable_on cube n" | 
| 149 | by (auto dest: lebesgueD) | |
| 46731 | 150 | let ?m = "\<lambda>n i. integral (cube n) (indicator (A i) :: _\<Rightarrow>real)" | 
| 151 | let ?M = "\<lambda>n I. integral (cube n) (indicator (\<Union>i\<in>I. A i) :: _\<Rightarrow>real)" | |
| 47694 | 152 | have nn[simp, intro]: "\<And>i n. 0 \<le> ?m n i" by (auto intro!: Integration.integral_nonneg) | 
| 41654 | 153 | assume "(\<Union>i. A i) \<in> sets lebesgue" | 
| 154 | then have UN_A[simp, intro]: "\<And>i n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n" | |
| 47694 | 155 | by (auto simp: sets_lebesgue) | 
| 156 | show "(\<Sum>n. ?\<mu> (A n)) = ?\<mu> (\<Union>i. A i)" | |
| 157 | proof (subst suminf_SUP_eq, safe intro!: incseq_SucI) | |
| 43920 | 158 | fix i n show "ereal (?m n i) \<le> ereal (?m (Suc n) i)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 159 | using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI) | 
| 41654 | 160 | next | 
| 43920 | 161 | fix i n show "0 \<le> ereal (?m n i)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 162 | using rA unfolding lebesgue_def | 
| 44928 
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
 hoelzl parents: 
44890diff
changeset | 163 | by (auto intro!: SUP_upper2 integral_nonneg) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 164 | next | 
| 43920 | 165 | show "(SUP n. \<Sum>i. ereal (?m n i)) = (SUP n. ereal (?M n UNIV))" | 
| 166 | proof (intro arg_cong[where f="SUPR UNIV"] ext sums_unique[symmetric] sums_ereal[THEN iffD2] sums_def[THEN iffD2]) | |
| 41654 | 167 | fix n | 
| 168 |         have "\<And>m. (UNION {..<m} A) \<in> sets lebesgue" using rA by auto
 | |
| 169 | from lebesgueD[OF this] | |
| 170 |         have "(\<lambda>m. ?M n {..< m}) ----> ?M n UNIV"
 | |
| 171 | (is "(\<lambda>m. integral _ (?A m)) ----> ?I") | |
| 172 | by (intro dominated_convergence(2)[where f="?A" and h="\<lambda>x. 1::real"]) | |
| 173 | (auto intro: LIMSEQ_indicator_UN simp: cube_def) | |
| 174 | moreover | |
| 175 |         { fix m have *: "(\<Sum>x<m. ?m n x) = ?M n {..< m}"
 | |
| 176 | proof (induct m) | |
| 177 | case (Suc m) | |
| 178 | have "(\<Union>i<m. A i) \<in> sets lebesgue" using rA by auto | |
| 179 | then have "(indicator (\<Union>i<m. A i) :: _\<Rightarrow>real) integrable_on (cube n)" | |
| 180 | by (auto dest!: lebesgueD) | |
| 181 | moreover | |
| 182 |             have "(\<Union>i<m. A i) \<inter> A m = {}"
 | |
| 183 | using rA(2)[unfolded disjoint_family_on_def, THEN bspec, of m] | |
| 184 | by auto | |
| 185 | then have "\<And>x. indicator (\<Union>i<Suc m. A i) x = | |
| 186 | indicator (\<Union>i<m. A i) x + (indicator (A m) x :: real)" | |
| 187 | by (auto simp: indicator_add lessThan_Suc ac_simps) | |
| 188 | ultimately show ?case | |
| 47694 | 189 | using Suc A by (simp add: Integration.integral_add[symmetric]) | 
| 41654 | 190 | qed auto } | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 191 | ultimately show "(\<lambda>m. \<Sum>x = 0..<m. ?m n x) ----> ?M n UNIV" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 192 | by (simp add: atLeast0LessThan) | 
| 41654 | 193 | qed | 
| 194 | qed | |
| 195 | qed | |
| 47694 | 196 | next | 
| 197 | qed (auto, fact) | |
| 40859 | 198 | |
| 41654 | 199 | lemma has_integral_interval_cube: | 
| 200 | fixes a b :: "'a::ordered_euclidean_space" | |
| 201 |   shows "(indicator {a .. b} has_integral
 | |
| 202 |     content ({\<chi>\<chi> i. max (- real n) (a $$ i) .. \<chi>\<chi> i. min (real n) (b $$ i)} :: 'a set)) (cube n)"
 | |
| 203 | (is "(?I has_integral content ?R) (cube n)") | |
| 40859 | 204 | proof - | 
| 41654 | 205 |   let "{?N .. ?P}" = ?R
 | 
| 206 | have [simp]: "(\<lambda>x. if x \<in> cube n then ?I x else 0) = indicator ?R" | |
| 207 | by (auto simp: indicator_def cube_def fun_eq_iff eucl_le[where 'a='a]) | |
| 208 | have "(?I has_integral content ?R) (cube n) \<longleftrightarrow> (indicator ?R has_integral content ?R) UNIV" | |
| 209 | unfolding has_integral_restrict_univ[where s="cube n", symmetric] by simp | |
| 210 | also have "\<dots> \<longleftrightarrow> ((\<lambda>x. 1) has_integral content ?R) ?R" | |
| 46905 | 211 | unfolding indicator_def [abs_def] has_integral_restrict_univ .. | 
| 41654 | 212 | finally show ?thesis | 
| 213 | using has_integral_const[of "1::real" "?N" "?P"] by simp | |
| 40859 | 214 | qed | 
| 38656 | 215 | |
| 41654 | 216 | lemma lebesgueI_borel[intro, simp]: | 
| 217 | fixes s::"'a::ordered_euclidean_space set" | |
| 40859 | 218 | assumes "s \<in> sets borel" shows "s \<in> sets lebesgue" | 
| 41654 | 219 | proof - | 
| 47694 | 220 |   have "s \<in> sigma_sets (space lebesgue) (range (\<lambda>(a, b). {a .. (b :: 'a\<Colon>ordered_euclidean_space)}))"
 | 
| 221 | using assms by (simp add: borel_eq_atLeastAtMost) | |
| 222 | also have "\<dots> \<subseteq> sets lebesgue" | |
| 223 | proof (safe intro!: sigma_sets_subset lebesgueI) | |
| 41654 | 224 | fix n :: nat and a b :: 'a | 
| 225 | let ?N = "\<chi>\<chi> i. max (- real n) (a $$ i)" | |
| 226 | let ?P = "\<chi>\<chi> i. min (real n) (b $$ i)" | |
| 227 |     show "(indicator {a..b} :: 'a\<Rightarrow>real) integrable_on cube n"
 | |
| 228 | unfolding integrable_on_def | |
| 229 | using has_integral_interval_cube[of a b] by auto | |
| 230 | qed | |
| 47694 | 231 | finally show ?thesis . | 
| 38656 | 232 | qed | 
| 233 | ||
| 47757 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 234 | lemma borel_measurable_lebesgueI: | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 235 | "f \<in> borel_measurable borel \<Longrightarrow> f \<in> borel_measurable lebesgue" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 236 | unfolding measurable_def by simp | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 237 | |
| 40859 | 238 | lemma lebesgueI_negligible[dest]: fixes s::"'a::ordered_euclidean_space set" | 
| 239 | assumes "negligible s" shows "s \<in> sets lebesgue" | |
| 41654 | 240 | using assms by (force simp: cube_def integrable_on_def negligible_def intro!: lebesgueI) | 
| 38656 | 241 | |
| 41654 | 242 | lemma lmeasure_eq_0: | 
| 47694 | 243 | fixes S :: "'a::ordered_euclidean_space set" | 
| 244 | assumes "negligible S" shows "emeasure lebesgue S = 0" | |
| 40859 | 245 | proof - | 
| 41654 | 246 | have "\<And>n. integral (cube n) (indicator S :: 'a\<Rightarrow>real) = 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 | 247 | unfolding lebesgue_integral_def using 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 | 248 | by (intro integral_unique some1_equality ex_ex1I) | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 249 | (auto simp: cube_def negligible_def) | 
| 47694 | 250 | then show ?thesis | 
| 251 | using assms by (simp add: emeasure_lebesgue lebesgueI_negligible) | |
| 40859 | 252 | qed | 
| 253 | ||
| 254 | lemma lmeasure_iff_LIMSEQ: | |
| 47694 | 255 | assumes A: "A \<in> sets lebesgue" and "0 \<le> m" | 
| 256 | shows "emeasure lebesgue A = ereal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m" | |
| 257 | proof (subst emeasure_lebesgue[OF A], intro SUP_eq_LIMSEQ) | |
| 41654 | 258 | show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))" | 
| 259 | using cube_subset assms by (intro monoI integral_subset_le) (auto dest!: lebesgueD) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 260 | qed | 
| 38656 | 261 | |
| 41654 | 262 | lemma has_integral_indicator_UNIV: | 
| 263 | fixes s A :: "'a::ordered_euclidean_space set" and x :: real | |
| 264 | shows "((indicator (s \<inter> A) :: 'a\<Rightarrow>real) has_integral x) UNIV = ((indicator s :: _\<Rightarrow>real) has_integral x) A" | |
| 265 | proof - | |
| 266 | have "(\<lambda>x. if x \<in> A then indicator s x else 0) = (indicator (s \<inter> A) :: _\<Rightarrow>real)" | |
| 267 | by (auto simp: fun_eq_iff indicator_def) | |
| 268 | then show ?thesis | |
| 269 | unfolding has_integral_restrict_univ[where s=A, symmetric] by simp | |
| 40859 | 270 | qed | 
| 38656 | 271 | |
| 41654 | 272 | lemma | 
| 273 | fixes s a :: "'a::ordered_euclidean_space set" | |
| 274 | shows integral_indicator_UNIV: | |
| 275 | "integral UNIV (indicator (s \<inter> A) :: 'a\<Rightarrow>real) = integral A (indicator s :: _\<Rightarrow>real)" | |
| 276 | and integrable_indicator_UNIV: | |
| 277 | "(indicator (s \<inter> A) :: 'a\<Rightarrow>real) integrable_on UNIV \<longleftrightarrow> (indicator s :: 'a\<Rightarrow>real) integrable_on A" | |
| 278 | unfolding integral_def integrable_on_def has_integral_indicator_UNIV by auto | |
| 279 | ||
| 280 | lemma lmeasure_finite_has_integral: | |
| 281 | fixes s :: "'a::ordered_euclidean_space set" | |
| 47694 | 282 | assumes "s \<in> sets lebesgue" "emeasure lebesgue s = ereal m" "0 \<le> m" | 
| 41654 | 283 | shows "(indicator s has_integral m) UNIV" | 
| 284 | proof - | |
| 285 | let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real" | |
| 286 | have **: "(?I s) integrable_on UNIV \<and> (\<lambda>k. integral UNIV (?I (s \<inter> cube k))) ----> integral UNIV (?I s)" | |
| 287 | proof (intro monotone_convergence_increasing allI ballI) | |
| 288 | have LIMSEQ: "(\<lambda>n. integral (cube n) (?I s)) ----> m" | |
| 289 | using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1, 3)] . | |
| 290 |     { fix n have "integral (cube n) (?I s) \<le> m"
 | |
| 291 | using cube_subset assms | |
| 292 | by (intro incseq_le[where L=m] LIMSEQ incseq_def[THEN iffD2] integral_subset_le allI impI) | |
| 293 | (auto dest!: lebesgueD) } | |
| 294 | moreover | |
| 295 |     { fix n have "0 \<le> integral (cube n) (?I s)"
 | |
| 47694 | 296 | using assms by (auto dest!: lebesgueD intro!: Integration.integral_nonneg) } | 
| 41654 | 297 | ultimately | 
| 298 |     show "bounded {integral UNIV (?I (s \<inter> cube k)) |k. True}"
 | |
| 299 | unfolding bounded_def | |
| 300 | apply (rule_tac exI[of _ 0]) | |
| 301 | apply (rule_tac exI[of _ m]) | |
| 302 | by (auto simp: dist_real_def integral_indicator_UNIV) | |
| 303 | fix k show "?I (s \<inter> cube k) integrable_on UNIV" | |
| 304 | unfolding integrable_indicator_UNIV using assms by (auto dest!: lebesgueD) | |
| 305 | fix x show "?I (s \<inter> cube k) x \<le> ?I (s \<inter> cube (Suc k)) x" | |
| 306 | using cube_subset[of k "Suc k"] by (auto simp: indicator_def) | |
| 307 | next | |
| 308 | fix x :: 'a | |
| 309 | from mem_big_cube obtain k where k: "x \<in> cube k" . | |
| 310 |     { fix n have "?I (s \<inter> cube (n + k)) x = ?I s x"
 | |
| 311 | using k cube_subset[of k "n + k"] by (auto simp: indicator_def) } | |
| 312 | note * = this | |
| 313 | show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x" | |
| 314 | by (rule LIMSEQ_offset[where k=k]) (auto simp: *) | |
| 315 | qed | |
| 316 | note ** = conjunctD2[OF this] | |
| 317 | have m: "m = integral UNIV (?I s)" | |
| 318 | apply (intro LIMSEQ_unique[OF _ **(2)]) | |
| 319 | using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1,3)] integral_indicator_UNIV . | |
| 320 | show ?thesis | |
| 321 | unfolding m by (intro integrable_integral **) | |
| 38656 | 322 | qed | 
| 323 | ||
| 47694 | 324 | lemma lmeasure_finite_integrable: assumes s: "s \<in> sets lebesgue" and "emeasure lebesgue s \<noteq> \<infinity>" | 
| 41654 | 325 | shows "(indicator s :: _ \<Rightarrow> real) integrable_on UNIV" | 
| 47694 | 326 | proof (cases "emeasure lebesgue s") | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 327 | case (real m) | 
| 47694 | 328 | with lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this] emeasure_nonneg[of lebesgue s] | 
| 41654 | 329 | show ?thesis unfolding integrable_on_def by auto | 
| 47694 | 330 | qed (insert assms emeasure_nonneg[of lebesgue s], auto) | 
| 38656 | 331 | |
| 41654 | 332 | lemma has_integral_lebesgue: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV" | 
| 333 | shows "s \<in> sets lebesgue" | |
| 334 | proof (intro lebesgueI) | |
| 335 | let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real" | |
| 336 | fix n show "(?I s) integrable_on cube n" unfolding cube_def | |
| 337 | proof (intro integrable_on_subinterval) | |
| 338 | show "(?I s) integrable_on UNIV" | |
| 339 | unfolding integrable_on_def using assms by auto | |
| 340 | qed auto | |
| 38656 | 341 | qed | 
| 342 | ||
| 41654 | 343 | lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV" | 
| 47694 | 344 | shows "emeasure lebesgue s = ereal m" | 
| 41654 | 345 | proof (intro lmeasure_iff_LIMSEQ[THEN iffD2]) | 
| 346 | let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real" | |
| 347 | show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] . | |
| 348 | show "0 \<le> m" using assms by (rule has_integral_nonneg) auto | |
| 349 | have "(\<lambda>n. integral UNIV (?I (s \<inter> cube n))) ----> integral UNIV (?I s)" | |
| 350 | proof (intro dominated_convergence(2) ballI) | |
| 351 | show "(?I s) integrable_on UNIV" unfolding integrable_on_def using assms by auto | |
| 352 | fix n show "?I (s \<inter> cube n) integrable_on UNIV" | |
| 353 | unfolding integrable_indicator_UNIV using `s \<in> sets lebesgue` by (auto dest: lebesgueD) | |
| 354 | fix x show "norm (?I (s \<inter> cube n) x) \<le> ?I s x" by (auto simp: indicator_def) | |
| 355 | next | |
| 356 | fix x :: 'a | |
| 357 | from mem_big_cube obtain k where k: "x \<in> cube k" . | |
| 358 |     { fix n have "?I (s \<inter> cube (n + k)) x = ?I s x"
 | |
| 359 | using k cube_subset[of k "n + k"] by (auto simp: indicator_def) } | |
| 360 | note * = this | |
| 361 | show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x" | |
| 362 | by (rule LIMSEQ_offset[where k=k]) (auto simp: *) | |
| 363 | qed | |
| 364 | then show "(\<lambda>n. integral (cube n) (?I s)) ----> m" | |
| 365 | unfolding integral_unique[OF assms] integral_indicator_UNIV by simp | |
| 366 | qed | |
| 367 | ||
| 368 | lemma has_integral_iff_lmeasure: | |
| 47694 | 369 | "(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> 0 \<le> m \<and> emeasure lebesgue A = ereal m)" | 
| 40859 | 370 | proof | 
| 41654 | 371 | assume "(indicator A has_integral m) UNIV" | 
| 372 | with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this] | |
| 47694 | 373 | show "A \<in> sets lebesgue \<and> 0 \<le> m \<and> emeasure lebesgue A = ereal m" | 
| 41654 | 374 | by (auto intro: has_integral_nonneg) | 
| 40859 | 375 | next | 
| 47694 | 376 | assume "A \<in> sets lebesgue \<and> 0 \<le> m \<and> emeasure lebesgue A = ereal m" | 
| 41654 | 377 | then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto | 
| 38656 | 378 | qed | 
| 379 | ||
| 41654 | 380 | lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" | 
| 47694 | 381 | shows "emeasure lebesgue s = ereal (integral UNIV (indicator s))" | 
| 41654 | 382 | using assms unfolding integrable_on_def | 
| 383 | proof safe | |
| 384 | fix y :: real assume "(indicator s has_integral y) UNIV" | |
| 385 | from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this] | |
| 47694 | 386 | show "emeasure lebesgue s = ereal (integral UNIV (indicator s))" by simp | 
| 40859 | 387 | qed | 
| 38656 | 388 | |
| 389 | lemma lebesgue_simple_function_indicator: | |
| 43920 | 390 | fixes f::"'a::ordered_euclidean_space \<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 | 391 | assumes f:"simple_function lebesgue f" | 
| 38656 | 392 |   shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))"
 | 
| 47694 | 393 | by (rule, subst simple_function_indicator_representation[OF f]) auto | 
| 38656 | 394 | |
| 41654 | 395 | lemma integral_eq_lmeasure: | 
| 47694 | 396 | "(indicator s::_\<Rightarrow>real) integrable_on UNIV \<Longrightarrow> integral UNIV (indicator s) = real (emeasure lebesgue s)" | 
| 41654 | 397 | by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg) | 
| 38656 | 398 | |
| 47694 | 399 | lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "emeasure lebesgue s \<noteq> \<infinity>" | 
| 41654 | 400 | using lmeasure_eq_integral[OF assms] by auto | 
| 38656 | 401 | |
| 40859 | 402 | lemma negligible_iff_lebesgue_null_sets: | 
| 47694 | 403 | "negligible A \<longleftrightarrow> A \<in> null_sets lebesgue" | 
| 40859 | 404 | proof | 
| 405 | assume "negligible A" | |
| 406 | from this[THEN lebesgueI_negligible] this[THEN lmeasure_eq_0] | |
| 47694 | 407 | show "A \<in> null_sets lebesgue" by auto | 
| 40859 | 408 | next | 
| 47694 | 409 | assume A: "A \<in> null_sets lebesgue" | 
| 410 | then have *:"((indicator A) has_integral (0::real)) UNIV" using lmeasure_finite_has_integral[of A] | |
| 411 | by (auto simp: null_sets_def) | |
| 41654 | 412 | show "negligible A" unfolding negligible_def | 
| 413 | proof (intro allI) | |
| 414 | fix a b :: 'a | |
| 415 |     have integrable: "(indicator A :: _\<Rightarrow>real) integrable_on {a..b}"
 | |
| 416 | by (intro integrable_on_subinterval has_integral_integrable) (auto intro: *) | |
| 417 |     then have "integral {a..b} (indicator A) \<le> (integral UNIV (indicator A) :: real)"
 | |
| 47694 | 418 | using * by (auto intro!: integral_subset_le) | 
| 41654 | 419 |     moreover have "(0::real) \<le> integral {a..b} (indicator A)"
 | 
| 420 | using integrable by (auto intro!: integral_nonneg) | |
| 421 |     ultimately have "integral {a..b} (indicator A) = (0::real)"
 | |
| 422 | using integral_unique[OF *] by auto | |
| 423 |     then show "(indicator A has_integral (0::real)) {a..b}"
 | |
| 424 | using integrable_integral[OF integrable] by simp | |
| 425 | qed | |
| 426 | qed | |
| 427 | ||
| 428 | lemma integral_const[simp]: | |
| 429 | fixes a b :: "'a::ordered_euclidean_space" | |
| 430 |   shows "integral {a .. b} (\<lambda>x. c) = content {a .. b} *\<^sub>R c"
 | |
| 431 | by (rule integral_unique) (rule has_integral_const) | |
| 432 | ||
| 47694 | 433 | lemma lmeasure_UNIV[intro]: "emeasure lebesgue (UNIV::'a::ordered_euclidean_space set) = \<infinity>" | 
| 434 | proof (simp add: emeasure_lebesgue, intro SUP_PInfty bexI) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 435 | fix n :: nat | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 436 | have "indicator UNIV = (\<lambda>x::'a. 1 :: real)" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 437 | moreover | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 438 |   { have "real n \<le> (2 * real n) ^ DIM('a)"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 439 | proof (cases n) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 440 | case 0 then show ?thesis by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 441 | next | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 442 | case (Suc n') | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 443 | have "real n \<le> (2 * real n)^1" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 444 |       also have "(2 * real n)^1 \<le> (2 * real n) ^ DIM('a)"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 445 | using Suc DIM_positive[where 'a='a] by (intro power_increasing) (auto simp: real_of_nat_Suc) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 446 | finally show ?thesis . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 447 | qed } | 
| 43920 | 448 | ultimately show "ereal (real n) \<le> ereal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 449 | using integral_const DIM_positive[where 'a='a] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 450 | by (auto simp: cube_def content_closed_interval_cases setprod_constant) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 451 | qed simp | 
| 40859 | 452 | |
| 453 | lemma | |
| 454 | fixes a b ::"'a::ordered_euclidean_space" | |
| 47694 | 455 |   shows lmeasure_atLeastAtMost[simp]: "emeasure lebesgue {a..b} = ereal (content {a..b})"
 | 
| 41654 | 456 | proof - | 
| 457 |   have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV"
 | |
| 46905 | 458 | unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def]) | 
| 41654 | 459 | from lmeasure_eq_integral[OF this] show ?thesis unfolding integral_indicator_UNIV | 
| 46905 | 460 | by (simp add: indicator_def [abs_def]) | 
| 40859 | 461 | qed | 
| 462 | ||
| 463 | lemma atLeastAtMost_singleton_euclidean[simp]: | |
| 464 |   fixes a :: "'a::ordered_euclidean_space" shows "{a .. a} = {a}"
 | |
| 465 | by (force simp: eucl_le[where 'a='a] euclidean_eq[where 'a='a]) | |
| 466 | ||
| 467 | lemma content_singleton[simp]: "content {a} = 0"
 | |
| 468 | proof - | |
| 469 |   have "content {a .. a} = 0"
 | |
| 470 | by (subst content_closed_interval) auto | |
| 471 | then show ?thesis by simp | |
| 472 | qed | |
| 473 | ||
| 474 | lemma lmeasure_singleton[simp]: | |
| 47694 | 475 |   fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0"
 | 
| 41654 | 476 | using lmeasure_atLeastAtMost[of a a] by simp | 
| 40859 | 477 | |
| 478 | declare content_real[simp] | |
| 479 | ||
| 480 | lemma | |
| 481 | fixes a b :: real | |
| 482 | shows lmeasure_real_greaterThanAtMost[simp]: | |
| 47694 | 483 |     "emeasure lebesgue {a <.. b} = ereal (if a \<le> b then b - a else 0)"
 | 
| 40859 | 484 | proof cases | 
| 485 | assume "a < b" | |
| 47694 | 486 |   then have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b} - emeasure lebesgue {a}"
 | 
| 487 | by (subst emeasure_Diff[symmetric]) | |
| 488 | (auto intro!: arg_cong[where f="emeasure lebesgue"]) | |
| 40859 | 489 | then show ?thesis by auto | 
| 490 | qed auto | |
| 491 | ||
| 492 | lemma | |
| 493 | fixes a b :: real | |
| 494 | shows lmeasure_real_atLeastLessThan[simp]: | |
| 47694 | 495 |     "emeasure lebesgue {a ..< b} = ereal (if a \<le> b then b - a else 0)"
 | 
| 40859 | 496 | proof cases | 
| 497 | assume "a < b" | |
| 47694 | 498 |   then have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b} - emeasure lebesgue {b}"
 | 
| 499 | by (subst emeasure_Diff[symmetric]) | |
| 500 | (auto intro!: arg_cong[where f="emeasure lebesgue"]) | |
| 41654 | 501 | then show ?thesis by auto | 
| 502 | qed auto | |
| 503 | ||
| 504 | lemma | |
| 505 | fixes a b :: real | |
| 506 | shows lmeasure_real_greaterThanLessThan[simp]: | |
| 47694 | 507 |     "emeasure lebesgue {a <..< b} = ereal (if a \<le> b then b - a else 0)"
 | 
| 41654 | 508 | proof cases | 
| 509 | assume "a < b" | |
| 47694 | 510 |   then have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a <.. b} - emeasure lebesgue {b}"
 | 
| 511 | by (subst emeasure_Diff[symmetric]) | |
| 512 | (auto intro!: arg_cong[where f="emeasure lebesgue"]) | |
| 40859 | 513 | then show ?thesis by auto | 
| 514 | qed auto | |
| 515 | ||
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 516 | subsection {* Lebesgue-Borel measure *}
 | 
| 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 517 | |
| 47694 | 518 | definition "lborel = measure_of UNIV (sets borel) (emeasure lebesgue)" | 
| 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 | 519 | |
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 520 | lemma | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 521 | shows space_lborel[simp]: "space lborel = UNIV" | 
| 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 hoelzl parents: 
41661diff
changeset | 522 | and sets_lborel[simp]: "sets lborel = sets borel" | 
| 47694 | 523 | and measurable_lborel1[simp]: "measurable lborel = measurable borel" | 
| 524 | and measurable_lborel2[simp]: "measurable A lborel = measurable A borel" | |
| 525 | using sigma_sets_eq[of borel] | |
| 526 | by (auto simp add: lborel_def measurable_def[abs_def]) | |
| 40859 | 527 | |
| 47694 | 528 | lemma emeasure_lborel[simp]: "A \<in> sets borel \<Longrightarrow> emeasure lborel A = emeasure lebesgue A" | 
| 529 | by (rule emeasure_measure_of[OF lborel_def]) | |
| 530 | (auto simp: positive_def emeasure_nonneg countably_additive_def intro!: suminf_emeasure) | |
| 40859 | 531 | |
| 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 | 532 | interpretation lborel: sigma_finite_measure lborel | 
| 47694 | 533 | proof (default, intro conjI exI[of _ "\<lambda>n. cube n"]) | 
| 534 | show "range cube \<subseteq> sets lborel" by (auto intro: borel_closed) | |
| 535 |   { fix x :: 'a have "\<exists>n. x\<in>cube n" using mem_big_cube by auto }
 | |
| 536 | then show "(\<Union>i. cube i) = (space lborel :: 'a set)" using mem_big_cube by auto | |
| 537 | show "\<forall>i. emeasure lborel (cube i) \<noteq> \<infinity>" by (simp add: cube_def) | |
| 538 | 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 | 539 | |
| 
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 | interpretation lebesgue: sigma_finite_measure lebesgue | 
| 40859 | 541 | proof | 
| 47694 | 542 | from lborel.sigma_finite guess A :: "nat \<Rightarrow> 'a set" .. | 
| 543 | then show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. emeasure lebesgue (A i) \<noteq> \<infinity>)" | |
| 544 | by (intro exI[of _ A]) (auto simp: subset_eq) | |
| 40859 | 545 | qed | 
| 546 | ||
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 547 | subsection {* Lebesgue integrable implies Gauge integrable *}
 | 
| 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 548 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 549 | lemma has_integral_cmult_real: | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 550 | fixes c :: real | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 551 | assumes "c \<noteq> 0 \<Longrightarrow> (f has_integral x) A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 552 | shows "((\<lambda>x. c * f x) has_integral c * x) A" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 553 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 554 | assume "c \<noteq> 0" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 555 | from has_integral_cmul[OF assms[OF this], of c] show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 556 | unfolding real_scaleR_def . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 557 | qed simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 558 | |
| 40859 | 559 | lemma simple_function_has_integral: | 
| 43920 | 560 | fixes f::"'a::ordered_euclidean_space \<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 | 561 | assumes f:"simple_function lebesgue f" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 562 |   and f':"range f \<subseteq> {0..<\<infinity>}"
 | 
| 47694 | 563 |   and om:"\<And>x. x \<in> range f \<Longrightarrow> emeasure lebesgue (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> 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 | 564 | shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 565 | unfolding simple_integral_def space_lebesgue | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 566 | proof (subst lebesgue_simple_function_indicator) | 
| 47694 | 567 |   let ?M = "\<lambda>x. emeasure lebesgue (f -` {x} \<inter> UNIV)"
 | 
| 46731 | 568 |   let ?F = "\<lambda>x. indicator (f -` {x})"
 | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 569 |   { fix x y assume "y \<in> range f"
 | 
| 43920 | 570 | from subsetD[OF f' this] have "y * ?F y x = ereal (real y * ?F y x)" | 
| 571 | by (cases rule: ereal2_cases[of y "?F y x"]) | |
| 572 | (auto simp: indicator_def one_ereal_def split: split_if_asm) } | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 573 | moreover | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 574 |   { fix x assume x: "x\<in>range f"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 575 | have "x * ?M x = real x * real (?M x)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 576 | proof cases | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 577 | assume "x \<noteq> 0" with om[OF x] have "?M x \<noteq> \<infinity>" by auto | 
| 47694 | 578 | with subsetD[OF f' x] f[THEN simple_functionD(2)] show ?thesis | 
| 43920 | 579 | by (cases rule: ereal2_cases[of x "?M x"]) auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 580 | qed simp } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 581 | ultimately | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 582 | have "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV \<longleftrightarrow> | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 583 | ((\<lambda>x. \<Sum>y\<in>range f. real y * ?F y x) has_integral (\<Sum>x\<in>range f. real x * real (?M x))) UNIV" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 584 | by simp | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 585 | also have \<dots> | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 586 | proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral | 
| 47694 | 587 | real_of_ereal_pos emeasure_nonneg ballI) | 
| 588 |     show *: "finite (range f)" "\<And>y. f -` {y} \<in> sets lebesgue"
 | |
| 589 | using simple_functionD[OF f] by auto | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 590 | fix y assume "real y \<noteq> 0" "y \<in> range f" | 
| 47694 | 591 |     with * om[OF this(2)] show "emeasure lebesgue (f -` {y}) = ereal (real (?M y))"
 | 
| 43920 | 592 | by (auto simp: ereal_real) | 
| 41654 | 593 | qed | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 594 | finally show "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV" . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 595 | qed fact | 
| 40859 | 596 | |
| 597 | lemma bounded_realI: assumes "\<forall>x\<in>s. abs (x::real) \<le> B" shows "bounded s" | |
| 598 | unfolding bounded_def dist_real_def apply(rule_tac x=0 in exI) | |
| 599 | using assms by auto | |
| 600 | ||
| 601 | lemma simple_function_has_integral': | |
| 43920 | 602 | fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 603 | assumes f: "simple_function lebesgue f" "\<And>x. 0 \<le> f x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 604 | and i: "integral\<^isup>S lebesgue f \<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 | 605 | shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 606 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 607 |   let ?f = "\<lambda>x. if x \<in> f -` {\<infinity>} then 0 else f x"
 | 
| 47694 | 608 | note f(1)[THEN simple_functionD(2)] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 609 | then have [simp, intro]: "\<And>X. f -` X \<in> sets lebesgue" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 610 | have f': "simple_function lebesgue ?f" | 
| 47694 | 611 | using f by (intro simple_function_If_set) auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 612 |   have rng: "range ?f \<subseteq> {0..<\<infinity>}" using f by auto
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 613 | have "AE x in lebesgue. f x = ?f x" | 
| 47694 | 614 | using simple_integral_PInf[OF f i] | 
| 615 |     by (intro AE_I[where N="f -` {\<infinity>} \<inter> space lebesgue"]) auto
 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 616 | from f(1) f' this have eq: "integral\<^isup>S lebesgue f = integral\<^isup>S lebesgue ?f" | 
| 47694 | 617 | by (rule simple_integral_cong_AE) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 618 | have real_eq: "\<And>x. real (f x) = real (?f x)" by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 619 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 620 | show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 621 | unfolding eq real_eq | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 622 | proof (rule simple_function_has_integral[OF f' rng]) | 
| 47694 | 623 |     fix x assume x: "x \<in> range ?f" and inf: "emeasure lebesgue (?f -` {x} \<inter> UNIV) = \<infinity>"
 | 
| 624 |     have "x * emeasure lebesgue (?f -` {x} \<inter> UNIV) = (\<integral>\<^isup>S y. x * indicator (?f -` {x}) y \<partial>lebesgue)"
 | |
| 625 | using f'[THEN simple_functionD(2)] | |
| 626 | by (simp add: simple_integral_cmult_indicator) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 627 | also have "\<dots> \<le> integral\<^isup>S lebesgue f" | 
| 47694 | 628 | using f'[THEN simple_functionD(2)] f | 
| 629 | by (intro simple_integral_mono simple_function_mult simple_function_indicator) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 630 | (auto split: split_indicator) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 631 | finally show "x = 0" unfolding inf using i subsetD[OF rng x] by (auto split: split_if_asm) | 
| 40859 | 632 | qed | 
| 633 | qed | |
| 634 | ||
| 635 | lemma positive_integral_has_integral: | |
| 43920 | 636 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 637 |   assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<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 | 638 | shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 639 | proof - | 
| 47694 | 640 | from borel_measurable_implies_simple_function_sequence'[OF f(1)] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 641 | guess u . note u = this | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 642 | have SUP_eq: "\<And>x. (SUP i. u i x) = f x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 643 | using u(4) f(2)[THEN subsetD] by (auto split: split_max) | 
| 46731 | 644 | let ?u = "\<lambda>i x. real (u i x)" | 
| 47694 | 645 | note u_eq = positive_integral_eq_simple_integral[OF u(1,5), symmetric] | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 646 |   { fix i
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 647 | note u_eq | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 648 | also have "integral\<^isup>P lebesgue (u i) \<le> (\<integral>\<^isup>+x. max 0 (f x) \<partial>lebesgue)" | 
| 47694 | 649 | by (intro positive_integral_mono) (auto intro: SUP_upper simp: u(4)[symmetric]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 650 | finally have "integral\<^isup>S lebesgue (u i) \<noteq> \<infinity>" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 651 | unfolding positive_integral_max_0 using f by auto } | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 652 | note u_fin = this | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 653 | then have u_int: "\<And>i. (?u i has_integral real (integral\<^isup>S lebesgue (u i))) UNIV" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 654 | by (rule simple_function_has_integral'[OF u(1,5)]) | 
| 43920 | 655 | have "\<forall>x. \<exists>r\<ge>0. f x = ereal r" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 656 | proof | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 657 | fix x from f(2) have "0 \<le> f x" "f x \<noteq> \<infinity>" by (auto simp: subset_eq) | 
| 43920 | 658 | then show "\<exists>r\<ge>0. f x = ereal r" by (cases "f x") auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 659 | qed | 
| 43920 | 660 | from choice[OF this] obtain f' where f': "f = (\<lambda>x. ereal (f' x))" "\<And>x. 0 \<le> f' x" by auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 661 | |
| 43920 | 662 | have "\<forall>i. \<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 663 | proof | 
| 43920 | 664 | fix i show "\<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 665 | proof (intro choice allI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 666 | fix i x have "u i x \<noteq> \<infinity>" using u(3)[of i] by (auto simp: image_iff) metis | 
| 43920 | 667 | then show "\<exists>r\<ge>0. u i x = ereal r" using u(5)[of i x] by (cases "u i x") auto | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 668 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 669 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 670 | from choice[OF this] obtain u' where | 
| 43920 | 671 | u': "u = (\<lambda>i x. ereal (u' i x))" "\<And>i x. 0 \<le> u' i x" by (auto simp: fun_eq_iff) | 
| 40859 | 672 | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 673 | have convergent: "f' integrable_on UNIV \<and> | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 674 | (\<lambda>k. integral UNIV (u' k)) ----> integral UNIV f'" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 675 | proof (intro monotone_convergence_increasing allI ballI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 676 | show int: "\<And>k. (u' k) integrable_on UNIV" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 677 | using u_int unfolding integrable_on_def u' by auto | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 678 | show "\<And>k x. u' k x \<le> u' (Suc k) x" using u(2,3,5) | 
| 43920 | 679 | by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_ereal_positive_mono) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 680 | show "\<And>x. (\<lambda>k. u' k x) ----> f' x" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 681 | using SUP_eq u(2) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 682 | by (intro SUP_eq_LIMSEQ[THEN iffD1]) (auto simp: u' f' incseq_mono incseq_Suc_iff le_fun_def) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 683 |     show "bounded {integral UNIV (u' k)|k. True}"
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 684 | proof (safe intro!: bounded_realI) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 685 | fix k | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 686 | have "\<bar>integral UNIV (u' k)\<bar> = integral UNIV (u' k)" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 687 | by (intro abs_of_nonneg integral_nonneg int ballI u') | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 688 | also have "\<dots> = real (integral\<^isup>S lebesgue (u k))" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 689 | using u_int[THEN integral_unique] by (simp add: u') | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 690 | also have "\<dots> = real (integral\<^isup>P lebesgue (u k))" | 
| 47694 | 691 | using positive_integral_eq_simple_integral[OF u(1,5)] by simp | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 692 | also have "\<dots> \<le> real (integral\<^isup>P lebesgue f)" using f | 
| 47694 | 693 | by (auto intro!: real_of_ereal_positive_mono positive_integral_positive | 
| 694 | positive_integral_mono SUP_upper simp: SUP_eq[symmetric]) | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 695 | finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^isup>P lebesgue f)" . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 696 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 697 | qed | 
| 40859 | 698 | |
| 43920 | 699 | have "integral\<^isup>P lebesgue f = ereal (integral UNIV f')" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 700 | proof (rule tendsto_unique[OF trivial_limit_sequentially]) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 701 | have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))" | 
| 47694 | 702 | unfolding u_eq by (intro LIMSEQ_ereal_SUPR incseq_positive_integral u) | 
| 703 | also note positive_integral_monotone_convergence_SUP | |
| 704 | [OF u(2) borel_measurable_simple_function[OF u(1)] u(5), symmetric] | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 705 | finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 706 | unfolding SUP_eq . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 707 | |
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 708 |     { fix k
 | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 709 | have "0 \<le> integral\<^isup>S lebesgue (u k)" | 
| 47694 | 710 | using u by (auto intro!: simple_integral_positive) | 
| 43920 | 711 | then have "integral\<^isup>S lebesgue (u k) = ereal (real (integral\<^isup>S lebesgue (u k)))" | 
| 712 | using u_fin by (auto simp: ereal_real) } | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 713 | note * = this | 
| 43920 | 714 | show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> ereal (integral UNIV f')" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 715 | using convergent using u_int[THEN integral_unique, symmetric] | 
| 47694 | 716 | by (subst *) (simp add: u') | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 717 | qed | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 718 | then show ?thesis using convergent by (simp add: f' integrable_integral) | 
| 40859 | 719 | qed | 
| 720 | ||
| 721 | lemma lebesgue_integral_has_integral: | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 722 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 723 | assumes f: "integrable lebesgue f" | 
| 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 | 724 | shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV" | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 725 | proof - | 
| 43920 | 726 | let ?n = "\<lambda>x. real (ereal (max 0 (- f x)))" and ?p = "\<lambda>x. real (ereal (max 0 (f x)))" | 
| 727 | have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: ereal_max) | |
| 47694 | 728 |   { fix f :: "'a \<Rightarrow> real" have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (max 0 (f x)) \<partial>lebesgue)"
 | 
| 729 | by (intro positive_integral_cong_pos) (auto split: split_max) } | |
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 730 | note eq = this | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 731 | show ?thesis | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 732 | unfolding lebesgue_integral_def | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 733 | apply (subst *) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 734 | apply (rule has_integral_sub) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 735 | unfolding eq[of f] eq[of "\<lambda>x. - f x"] | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 736 | apply (safe intro!: positive_integral_has_integral) | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 737 | using integrableD[OF f] | 
| 43920 | 738 | by (auto simp: zero_ereal_def[symmetric] positive_integral_max_0 split: split_max | 
| 47694 | 739 | intro!: measurable_If) | 
| 40859 | 740 | qed | 
| 741 | ||
| 47757 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 742 | lemma lebesgue_simple_integral_eq_borel: | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 743 | assumes f: "f \<in> borel_measurable borel" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 744 | shows "integral\<^isup>S lebesgue f = integral\<^isup>S lborel f" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 745 | using f[THEN measurable_sets] | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 746 | by (auto intro!: setsum_cong arg_cong2[where f="op *"] emeasure_lborel[symmetric] | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 747 | simp: simple_integral_def) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 748 | |
| 41546 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 749 | lemma lebesgue_positive_integral_eq_borel: | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 750 | assumes f: "f \<in> borel_measurable borel" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 751 | shows "integral\<^isup>P lebesgue f = integral\<^isup>P lborel f" | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 752 | proof - | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 753 | from f have "integral\<^isup>P lebesgue (\<lambda>x. max 0 (f x)) = integral\<^isup>P lborel (\<lambda>x. max 0 (f x))" | 
| 47694 | 754 | by (auto intro!: positive_integral_subalgebra[symmetric]) | 
| 41981 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 755 | then show ?thesis unfolding positive_integral_max_0 . | 
| 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 hoelzl parents: 
41831diff
changeset | 756 | qed | 
| 41546 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 757 | |
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 758 | lemma lebesgue_integral_eq_borel: | 
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 759 | assumes "f \<in> borel_measurable borel" | 
| 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 | 760 | shows "integrable lebesgue f \<longleftrightarrow> integrable lborel 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 | 761 | and "integral\<^isup>L lebesgue f = integral\<^isup>L lborel f" (is ?I) | 
| 41546 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 762 | 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 | 763 | have "sets lborel \<subseteq> sets lebesgue" by auto | 
| 47694 | 764 | from integral_subalgebra[of f lborel, OF _ this _ _] assms | 
| 41546 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 765 | show ?P ?I by auto | 
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 766 | qed | 
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 767 | |
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 768 | lemma borel_integral_has_integral: | 
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 769 | fixes f::"'a::ordered_euclidean_space => real" | 
| 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 | 770 | assumes f:"integrable lborel 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 | 771 | shows "(f has_integral (integral\<^isup>L lborel f)) UNIV" | 
| 41546 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 772 | proof - | 
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 773 | have borel: "f \<in> borel_measurable borel" | 
| 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 | 774 | using f unfolding integrable_def by auto | 
| 41546 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 775 | from f show ?thesis | 
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 776 | using lebesgue_integral_has_integral[of f] | 
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 777 | unfolding lebesgue_integral_eq_borel[OF borel] by simp | 
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 778 | qed | 
| 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 hoelzl parents: 
41097diff
changeset | 779 | |
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 780 | subsection {* Equivalence between product spaces and euclidean spaces *}
 | 
| 40859 | 781 | |
| 782 | definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> (nat \<Rightarrow> real)" where | |
| 783 |   "e2p x = (\<lambda>i\<in>{..<DIM('a)}. x$$i)"
 | |
| 784 | ||
| 785 | definition p2e :: "(nat \<Rightarrow> real) \<Rightarrow> 'a::ordered_euclidean_space" where | |
| 786 | "p2e x = (\<chi>\<chi> i. x i)" | |
| 787 | ||
| 41095 | 788 | lemma e2p_p2e[simp]: | 
| 789 |   "x \<in> extensional {..<DIM('a)} \<Longrightarrow> e2p (p2e x::'a::ordered_euclidean_space) = x"
 | |
| 790 | by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def) | |
| 40859 | 791 | |
| 41095 | 792 | lemma p2e_e2p[simp]: | 
| 793 | "p2e (e2p x) = (x::'a::ordered_euclidean_space)" | |
| 794 | by (auto simp: euclidean_eq[where 'a='a] p2e_def e2p_def) | |
| 40859 | 795 | |
| 47694 | 796 | interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure" | 
| 40859 | 797 | by default | 
| 798 | ||
| 47694 | 799 | interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "{..<n}" for n :: nat
 | 
| 800 | by default auto | |
| 801 | ||
| 802 | lemma bchoice_iff: "(\<forall>x\<in>A. \<exists>y. P x y) \<longleftrightarrow> (\<exists>f. \<forall>x\<in>A. P x (f x))" | |
| 803 | by metis | |
| 40859 | 804 | |
| 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 | 805 | lemma sets_product_borel: | 
| 47694 | 806 | assumes I: "finite I" | 
| 807 |   shows "sets (\<Pi>\<^isub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^isub>E i\<in>I. UNIV) { \<Pi>\<^isub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G")
 | |
| 808 | proof (subst sigma_prod_algebra_sigma_eq[where S="\<lambda>_ i::nat. {..<real i}" and E="\<lambda>_. range lessThan", OF I])
 | |
| 809 |   show "sigma_sets (space (Pi\<^isub>M I (\<lambda>i. lborel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G"
 | |
| 810 | by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff) | |
| 811 | qed (auto simp: borel_eq_lessThan incseq_def reals_Archimedean2 image_iff intro: real_natceiling_ge) | |
| 40859 | 812 | |
| 41661 | 813 | lemma measurable_e2p: | 
| 47694 | 814 |   "e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))"
 | 
| 815 | proof (rule measurable_sigma_sets[OF sets_product_borel]) | |
| 816 |   fix A :: "(nat \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i} |x. True} "
 | |
| 817 |   then obtain x where  "A = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {..<x i})" by auto
 | |
| 818 |   then have "e2p -` A = {..< (\<chi>\<chi> i. x i) :: 'a}"
 | |
| 819 | using DIM_positive by (auto simp add: Pi_iff set_eq_iff e2p_def | |
| 820 | euclidean_eq[where 'a='a] eucl_less[where 'a='a]) | |
| 821 | then show "e2p -` A \<inter> space (borel::'a measure) \<in> sets borel" by simp | |
| 822 | qed (auto simp: e2p_def) | |
| 41661 | 823 | |
| 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 | 824 | lemma measurable_p2e: | 
| 47694 | 825 |   "p2e \<in> measurable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. (lborel :: real measure))
 | 
| 826 | (borel :: 'a::ordered_euclidean_space measure)" | |
| 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 | 827 | (is "p2e \<in> measurable ?P _") | 
| 47694 | 828 | proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2]) | 
| 829 | fix x i | |
| 830 |   let ?A = "{w \<in> space ?P. (p2e w :: 'a) $$ i \<le> x}"
 | |
| 831 |   assume "i < DIM('a)"
 | |
| 832 |   then have "?A = (\<Pi>\<^isub>E j\<in>{..<DIM('a)}. if i = j then {.. x} else UNIV)"
 | |
| 833 | using DIM_positive by (auto simp: space_PiM p2e_def split: split_if_asm) | |
| 834 | then show "?A \<in> sets ?P" | |
| 835 | by auto | |
| 836 | qed | |
| 837 | ||
| 838 | lemma Int_stable_atLeastAtMost: | |
| 839 | fixes x::"'a::ordered_euclidean_space" | |
| 840 |   shows "Int_stable (range (\<lambda>(a, b::'a). {a..b}))"
 | |
| 841 | by (auto simp: inter_interval Int_stable_def) | |
| 41095 | 842 | |
| 47694 | 843 | lemma lborel_eqI: | 
| 844 | fixes M :: "'a::ordered_euclidean_space measure" | |
| 845 |   assumes emeasure_eq: "\<And>a b. emeasure M {a .. b} = content {a .. b}"
 | |
| 846 | assumes sets_eq: "sets M = sets borel" | |
| 847 | shows "lborel = M" | |
| 848 | proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost]) | |
| 849 |   let ?P = "\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel"
 | |
| 850 |   let ?E = "range (\<lambda>(a, b). {a..b} :: 'a set)"
 | |
| 851 | show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E" | |
| 852 | by (simp_all add: borel_eq_atLeastAtMost sets_eq) | |
| 47757 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 853 | |
| 47694 | 854 | show "range cube \<subseteq> ?E" unfolding cube_def [abs_def] by auto | 
| 855 | show "incseq cube" using cube_subset_Suc by (auto intro!: incseq_SucI) | |
| 856 |   { fix x :: 'a have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastforce }
 | |
| 857 | then show "(\<Union>i. cube i :: 'a set) = UNIV" by auto | |
| 858 | ||
| 859 |   { fix i show "emeasure lborel (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
 | |
| 860 |   { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
 | |
| 861 | by (auto simp: emeasure_eq) } | |
| 862 | qed | |
| 40859 | 863 | |
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 864 | lemma lborel_eq_lborel_space: | 
| 47694 | 865 |   "(lborel :: 'a measure) = distr (\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel) lborel p2e"
 | 
| 866 | (is "?B = ?D") | |
| 867 | proof (rule lborel_eqI) | |
| 868 | show "sets ?D = sets borel" by simp | |
| 869 |   let ?P = "(\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel)"
 | |
| 870 | fix a b :: 'a | |
| 871 |   have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^isub>E i\<in>{..<DIM('a)}. {a $$ i .. b $$ i})"
 | |
| 872 | by (auto simp: Pi_iff eucl_le[where 'a='a] p2e_def space_PiM) | |
| 873 |   have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}"
 | |
| 874 | proof cases | |
| 875 |     assume "{a..b} \<noteq> {}"
 | |
| 876 | then have "a \<le> b" | |
| 877 | by (simp add: interval_ne_empty eucl_le[where 'a='a]) | |
| 878 |     then have "emeasure lborel {a..b} = (\<Prod>x<DIM('a). emeasure lborel {a $$ x .. b $$ x})"
 | |
| 879 | by (auto simp: content_closed_interval eucl_le[where 'a='a] | |
| 880 | intro!: setprod_ereal[symmetric]) | |
| 881 |     also have "\<dots> = emeasure ?P (p2e -` {a..b} \<inter> space ?P)"
 | |
| 882 | unfolding * by (subst lborel_space.measure_times) auto | |
| 883 | finally show ?thesis by simp | |
| 884 | qed simp | |
| 885 |   then show "emeasure ?D {a .. b} = content {a .. b}"
 | |
| 886 | by (simp add: emeasure_distr measurable_p2e) | |
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 887 | qed | 
| 40859 | 888 | |
| 889 | lemma borel_fubini_positiv_integral: | |
| 43920 | 890 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal" | 
| 40859 | 891 | assumes f: "f \<in> borel_measurable borel" | 
| 47694 | 892 |   shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel)"
 | 
| 893 | by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f) | |
| 40859 | 894 | |
| 41704 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 895 | lemma borel_fubini_integrable: | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 896 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real" | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 897 | shows "integrable lborel f \<longleftrightarrow> | 
| 47694 | 898 |     integrable (\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel) (\<lambda>x. f (p2e x))"
 | 
| 41704 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 899 | (is "_ \<longleftrightarrow> integrable ?B ?f") | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 900 | proof | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 901 | assume "integrable lborel f" | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 902 | moreover then have f: "f \<in> borel_measurable borel" | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 903 | by auto | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 904 | moreover with measurable_p2e | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 905 | have "f \<circ> p2e \<in> borel_measurable ?B" | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 906 | by (rule measurable_comp) | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 907 | ultimately show "integrable ?B ?f" | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 908 | by (simp add: comp_def borel_fubini_positiv_integral integrable_def) | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 909 | next | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 910 | assume "integrable ?B ?f" | 
| 47694 | 911 | moreover | 
| 912 | then have "?f \<circ> e2p \<in> borel_measurable (borel::'a measure)" | |
| 913 | by (auto intro!: measurable_e2p) | |
| 41704 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 914 | then have "f \<in> borel_measurable borel" | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 915 | by (simp cong: measurable_cong) | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 916 | ultimately show "integrable lborel f" | 
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 917 | by (simp add: borel_fubini_positiv_integral integrable_def) | 
| 41704 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 918 | qed | 
| 
8c539202f854
add borel_fubini_integrable; remove unused bijectivity rules for measureable functions
 hoelzl parents: 
41689diff
changeset | 919 | |
| 40859 | 920 | lemma borel_fubini: | 
| 921 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real" | |
| 922 | assumes f: "f \<in> borel_measurable borel" | |
| 47694 | 923 |   shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^isub>M i\<in>{..<DIM('a)}. lborel))"
 | 
| 41706 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 hoelzl parents: 
41704diff
changeset | 924 | using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def) | 
| 38656 | 925 | |
| 47694 | 926 | lemma borel_measurable_indicator': | 
| 927 | "A \<in> sets borel \<Longrightarrow> f \<in> borel_measurable M \<Longrightarrow> (\<lambda>x. indicator A (f x)) \<in> borel_measurable M" | |
| 928 | using measurable_comp[OF _ borel_measurable_indicator, of f M borel A] by (auto simp add: comp_def) | |
| 42164 | 929 | |
| 930 | lemma lebesgue_real_affine: | |
| 47694 | 931 | fixes c :: real assumes "c \<noteq> 0" | 
| 932 | shows "lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. \<bar>c\<bar>)" (is "_ = ?D") | |
| 933 | proof (rule lborel_eqI) | |
| 934 |   fix a b show "emeasure ?D {a..b} = content {a .. b}"
 | |
| 935 | proof cases | |
| 936 | assume "0 < c" | |
| 937 |     then have "(\<lambda>x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}"
 | |
| 938 | by (auto simp: field_simps) | |
| 939 | with `0 < c` show ?thesis | |
| 940 | by (cases "a \<le> b") | |
| 941 | (auto simp: field_simps emeasure_density positive_integral_distr positive_integral_cmult | |
| 942 | borel_measurable_indicator' emeasure_distr) | |
| 943 | next | |
| 944 | assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto | |
| 945 |     then have *: "(\<lambda>x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}"
 | |
| 946 | by (auto simp: field_simps) | |
| 947 | with `c < 0` show ?thesis | |
| 948 | by (cases "a \<le> b") | |
| 949 | (auto simp: field_simps emeasure_density positive_integral_distr | |
| 950 | positive_integral_cmult borel_measurable_indicator' emeasure_distr) | |
| 42164 | 951 | qed | 
| 47694 | 952 | qed simp | 
| 42164 | 953 | |
| 47757 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 954 | lemma borel_cube[intro]: "cube n \<in> sets borel" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 955 | unfolding cube_def by auto | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 956 | |
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 957 | lemma integrable_on_cmult_iff: | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 958 | fixes c :: real assumes "c \<noteq> 0" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 959 | shows "(\<lambda>x. c * f x) integrable_on s \<longleftrightarrow> f integrable_on s" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 960 | using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c] `c \<noteq> 0` | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 961 | by auto | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 962 | |
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 963 | lemma positive_integral_borel_has_integral: | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 964 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 965 | assumes f_borel: "f \<in> borel_measurable borel" and nonneg: "\<And>x. 0 \<le> f x" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 966 | assumes I: "(f has_integral I) UNIV" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 967 | shows "(\<integral>\<^isup>+x. f x \<partial>lborel) = I" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 968 | proof - | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 969 | from f_borel have "(\<lambda>x. ereal (f x)) \<in> borel_measurable borel" by auto | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 970 | from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 971 | |
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 972 | have lebesgue_eq: "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (f x) \<partial>lborel)" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 973 | using f_borel by (intro lebesgue_positive_integral_eq_borel) auto | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 974 | also have "\<dots> = (SUP i. integral\<^isup>S lborel (F i))" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 975 | using F | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 976 | by (subst positive_integral_monotone_convergence_simple) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 977 | (simp_all add: positive_integral_max_0 simple_function_def) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 978 | also have "\<dots> \<le> ereal I" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 979 | proof (rule SUP_least) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 980 | fix i :: nat | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 981 | |
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 982 |     { fix z
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 983 | from F(4)[of z] have "F i z \<le> ereal (f z)" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 984 | by (metis SUP_upper UNIV_I ereal_max_0 min_max.sup_absorb2 nonneg) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 985 | with F(5)[of i z] have "real (F i z) \<le> f z" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 986 | by (cases "F i z") simp_all } | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 987 | note F_bound = this | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 988 | |
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 989 |     { fix x :: ereal assume x: "x \<noteq> 0" "x \<in> range (F i)"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 990 | with F(3,5)[of i] have [simp]: "real x \<noteq> 0" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 991 | by (metis image_iff order_eq_iff real_of_ereal_le_0) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 992 |       let ?s = "(\<lambda>n z. real x * indicator (F i -` {x} \<inter> cube n) z) :: nat \<Rightarrow> 'a \<Rightarrow> real"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 993 |       have "(\<lambda>z::'a. real x * indicator (F i -` {x}) z) integrable_on UNIV"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 994 | proof (rule dominated_convergence(1)) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 995 | fix n :: nat | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 996 |         have "(\<lambda>z. indicator (F i -` {x} \<inter> cube n) z :: real) integrable_on cube n"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 997 | using x F(1)[of i] | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 998 | by (intro lebesgueD) (auto simp: simple_function_def) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 999 | then have cube: "?s n integrable_on cube n" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1000 | by (simp add: integrable_on_cmult_iff) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1001 | show "?s n integrable_on UNIV" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1002 | by (rule integrable_on_superset[OF _ _ cube]) auto | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1003 | next | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1004 | show "f integrable_on UNIV" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1005 | unfolding integrable_on_def using I by auto | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1006 | next | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1007 | fix n from F_bound show "\<forall>x\<in>UNIV. norm (?s n x) \<le> f x" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1008 | using nonneg F(5) by (auto split: split_indicator) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1009 | next | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1010 |         show "\<forall>z\<in>UNIV. (\<lambda>n. ?s n z) ----> real x * indicator (F i -` {x}) z"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1011 | proof | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1012 | fix z :: 'a | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1013 | from mem_big_cube[of z] guess j . | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1014 |           then have x: "eventually (\<lambda>n. ?s n z = real x * indicator (F i -` {x}) z) sequentially"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1015 | by (auto intro!: eventually_sequentiallyI[where c=j] dest!: cube_subset split: split_indicator) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1016 |           then show "(\<lambda>n. ?s n z) ----> real x * indicator (F i -` {x}) z"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1017 | by (rule Lim_eventually) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1018 | qed | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1019 | qed | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1020 |       then have "(indicator (F i -` {x}) :: 'a \<Rightarrow> real) integrable_on UNIV"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1021 | by (simp add: integrable_on_cmult_iff) } | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1022 | note F_finite = lmeasure_finite[OF this] | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1023 | |
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1024 | have "((\<lambda>x. real (F i x)) has_integral real (integral\<^isup>S lebesgue (F i))) UNIV" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1025 | proof (rule simple_function_has_integral[of "F i"]) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1026 | show "simple_function lebesgue (F i)" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1027 | using F(1) by (simp add: simple_function_def) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1028 |       show "range (F i) \<subseteq> {0..<\<infinity>}"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1029 | using F(3,5)[of i] by (auto simp: image_iff) metis | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1030 | next | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1031 |       fix x assume "x \<in> range (F i)" "emeasure lebesgue (F i -` {x} \<inter> UNIV) = \<infinity>"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1032 | with F_finite[of x] show "x = 0" by auto | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1033 | qed | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1034 | from this I have "real (integral\<^isup>S lebesgue (F i)) \<le> I" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1035 | by (rule has_integral_le) (intro ballI F_bound) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1036 | moreover | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1037 |     { fix x assume x: "x \<in> range (F i)"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1038 | with F(3,5)[of i] have "x = 0 \<or> (0 < x \<and> x \<noteq> \<infinity>)" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1039 | by (auto simp: image_iff le_less) metis | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1040 |       with F_finite[OF _ x] x have "x * emeasure lebesgue (F i -` {x} \<inter> UNIV) \<noteq> \<infinity>"
 | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1041 | by auto } | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1042 | then have "integral\<^isup>S lebesgue (F i) \<noteq> \<infinity>" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1043 | unfolding simple_integral_def setsum_Pinfty space_lebesgue by blast | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1044 | moreover have "0 \<le> integral\<^isup>S lebesgue (F i)" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1045 | using F(1,5) by (intro simple_integral_positive) (auto simp: simple_function_def) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1046 | moreover have "integral\<^isup>S lebesgue (F i) = integral\<^isup>S lborel (F i)" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1047 | using F(1)[of i, THEN borel_measurable_simple_function] | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1048 | by (rule lebesgue_simple_integral_eq_borel) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1049 | ultimately show "integral\<^isup>S lborel (F i) \<le> ereal I" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1050 | by (cases "integral\<^isup>S lborel (F i)") auto | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1051 | qed | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1052 | also have "\<dots> < \<infinity>" by simp | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1053 | finally have finite: "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) \<noteq> \<infinity>" by simp | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1054 | have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable lebesgue" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1055 | using f_borel by (auto intro: borel_measurable_lebesgueI) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1056 | from positive_integral_has_integral[OF borel _ finite] | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1057 | have "(f has_integral real (\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue)) UNIV" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1058 | using nonneg by (simp add: subset_eq) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1059 | with I have "I = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue)" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1060 | by (rule has_integral_unique) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1061 | with finite positive_integral_positive[of _ "\<lambda>x. ereal (f x)"] show ?thesis | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1062 | by (cases "\<integral>\<^isup>+ x. ereal (f x) \<partial>lborel") (auto simp: lebesgue_eq) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1063 | qed | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1064 | |
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1065 | lemma has_integral_iff_positive_integral: | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1066 | fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1067 | assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1068 | shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^isup>P lborel f = I" | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1069 | using f positive_integral_borel_has_integral[of f I] positive_integral_has_integral[of f] | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1070 | by (auto simp: subset_eq borel_measurable_lebesgueI lebesgue_positive_integral_eq_borel) | 
| 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 hoelzl parents: 
47694diff
changeset | 1071 | |
| 38656 | 1072 | end |