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