author | hoelzl |
Wed, 18 Jun 2014 07:31:12 +0200 | |
changeset 57275 | 0ddb5b755cdc |
parent 57235 | b0b9a10e4bf4 |
child 57447 | 87429bdecad5 |
permissions | -rw-r--r-- |
42067 | 1 |
(* Title: HOL/Probability/Lebesgue_Measure.thy |
2 |
Author: Johannes Hölzl, TU München |
|
3 |
Author: Robert Himmelmann, TU München |
|
4 |
*) |
|
5 |
||
38656 | 6 |
header {* Lebsegue measure *} |
42067 | 7 |
|
38656 | 8 |
theory Lebesgue_Measure |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
9 |
imports Finite_Product_Measure Bochner_Integration |
38656 | 10 |
begin |
11 |
||
50104 | 12 |
lemma absolutely_integrable_on_indicator[simp]: |
13 |
fixes A :: "'a::ordered_euclidean_space set" |
|
14 |
shows "((indicator A :: _ \<Rightarrow> real) absolutely_integrable_on X) \<longleftrightarrow> |
|
15 |
(indicator A :: _ \<Rightarrow> real) integrable_on X" |
|
16 |
unfolding absolutely_integrable_on_def by simp |
|
49777 | 17 |
|
50104 | 18 |
lemma has_integral_indicator_UNIV: |
19 |
fixes s A :: "'a::ordered_euclidean_space set" and x :: real |
|
20 |
shows "((indicator (s \<inter> A) :: 'a\<Rightarrow>real) has_integral x) UNIV = ((indicator s :: _\<Rightarrow>real) has_integral x) A" |
|
21 |
proof - |
|
22 |
have "(\<lambda>x. if x \<in> A then indicator s x else 0) = (indicator (s \<inter> A) :: _\<Rightarrow>real)" |
|
23 |
by (auto simp: fun_eq_iff indicator_def) |
|
24 |
then show ?thesis |
|
25 |
unfolding has_integral_restrict_univ[where s=A, symmetric] by simp |
|
26 |
qed |
|
27 |
||
28 |
lemma |
|
29 |
fixes s a :: "'a::ordered_euclidean_space set" |
|
30 |
shows integral_indicator_UNIV: |
|
31 |
"integral UNIV (indicator (s \<inter> A) :: 'a\<Rightarrow>real) = integral A (indicator s :: _\<Rightarrow>real)" |
|
32 |
and integrable_indicator_UNIV: |
|
33 |
"(indicator (s \<inter> A) :: 'a\<Rightarrow>real) integrable_on UNIV \<longleftrightarrow> (indicator s :: 'a\<Rightarrow>real) integrable_on A" |
|
34 |
unfolding integral_def integrable_on_def has_integral_indicator_UNIV by auto |
|
47694 | 35 |
|
38656 | 36 |
subsection {* Standard Cubes *} |
37 |
||
40859 | 38 |
definition cube :: "nat \<Rightarrow> 'a::ordered_euclidean_space set" where |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
39 |
"cube n \<equiv> {\<Sum>i\<in>Basis. - n *\<^sub>R i .. \<Sum>i\<in>Basis. n *\<^sub>R i}" |
40859 | 40 |
|
49777 | 41 |
lemma borel_cube[intro]: "cube n \<in> sets borel" |
42 |
unfolding cube_def by auto |
|
43 |
||
40859 | 44 |
lemma cube_closed[intro]: "closed (cube n)" |
45 |
unfolding cube_def by auto |
|
46 |
||
47 |
lemma cube_subset[intro]: "n \<le> N \<Longrightarrow> cube n \<subseteq> cube N" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
48 |
by (fastforce simp: eucl_le[where 'a='a] cube_def setsum_negf) |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
49 |
|
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
50 |
lemma cube_subset_iff: "cube n \<subseteq> cube N \<longleftrightarrow> n \<le> N" |
56188 | 51 |
unfolding cube_def subset_box by (simp add: setsum_negf ex_in_conv eucl_le[where 'a='a]) |
38656 | 52 |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
53 |
lemma ball_subset_cube: "ball (0::'a::ordered_euclidean_space) (real n) \<subseteq> cube n" |
56188 | 54 |
apply (simp add: cube_def subset_eq mem_box setsum_negf eucl_le[where 'a='a]) |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
55 |
proof safe |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
56 |
fix x i :: 'a assume x: "x \<in> ball 0 (real n)" and i: "i \<in> Basis" |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
57 |
thus "- real n \<le> x \<bullet> i" "real n \<ge> x \<bullet> i" |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
58 |
using Basis_le_norm[OF i, of x] by(auto simp: dist_norm) |
38656 | 59 |
qed |
60 |
||
61 |
lemma mem_big_cube: obtains n where "x \<in> cube n" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
62 |
proof - |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
63 |
from reals_Archimedean2[of "norm x"] guess n .. |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
64 |
with ball_subset_cube[unfolded subset_eq, of n] |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
65 |
show ?thesis |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
66 |
by (intro that[where n=n]) (auto simp add: dist_norm) |
38656 | 67 |
qed |
68 |
||
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
|
69 |
lemma cube_subset_Suc[intro]: "cube n \<subseteq> cube (Suc n)" |
56188 | 70 |
unfolding cube_def cbox_interval[symmetric] subset_box by (simp add: setsum_negf) |
41654 | 71 |
|
50104 | 72 |
lemma has_integral_interval_cube: |
73 |
fixes a b :: "'a::ordered_euclidean_space" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
74 |
shows "(indicator {a .. b} has_integral content ({a .. b} \<inter> cube n)) (cube n)" |
50104 | 75 |
(is "(?I has_integral content ?R) (cube n)") |
76 |
proof - |
|
77 |
have [simp]: "(\<lambda>x. if x \<in> cube n then ?I x else 0) = indicator ?R" |
|
78 |
by (auto simp: indicator_def cube_def fun_eq_iff eucl_le[where 'a='a]) |
|
79 |
have "(?I has_integral content ?R) (cube n) \<longleftrightarrow> (indicator ?R has_integral content ?R) UNIV" |
|
80 |
unfolding has_integral_restrict_univ[where s="cube n", symmetric] by simp |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
81 |
also have "\<dots> \<longleftrightarrow> ((\<lambda>x. 1::real) has_integral content ?R *\<^sub>R 1) ?R" |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
82 |
unfolding indicator_def [abs_def] has_integral_restrict_univ real_scaleR_def mult_1_right .. |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
83 |
also have "((\<lambda>x. 1) has_integral content ?R *\<^sub>R 1) ?R" |
56188 | 84 |
unfolding cube_def inter_interval cbox_interval[symmetric] by (rule has_integral_const) |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
85 |
finally show ?thesis . |
50104 | 86 |
qed |
87 |
||
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
88 |
subsection {* Lebesgue measure *} |
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41704
diff
changeset
|
89 |
|
47694 | 90 |
definition lebesgue :: "'a::ordered_euclidean_space measure" where |
91 |
"lebesgue = measure_of UNIV {A. \<forall>n. (indicator A :: 'a \<Rightarrow> real) integrable_on cube n} |
|
92 |
(\<lambda>A. SUP n. ereal (integral (cube n) (indicator A)))" |
|
41661 | 93 |
|
41654 | 94 |
lemma space_lebesgue[simp]: "space lebesgue = UNIV" |
95 |
unfolding lebesgue_def by simp |
|
96 |
||
97 |
lemma lebesgueI: "(\<And>n. (indicator A :: _ \<Rightarrow> real) integrable_on cube n) \<Longrightarrow> A \<in> sets lebesgue" |
|
98 |
unfolding lebesgue_def by simp |
|
99 |
||
47694 | 100 |
lemma sigma_algebra_lebesgue: |
101 |
defines "leb \<equiv> {A. \<forall>n. (indicator A :: 'a::ordered_euclidean_space \<Rightarrow> real) integrable_on cube n}" |
|
102 |
shows "sigma_algebra UNIV leb" |
|
103 |
proof (safe intro!: sigma_algebra_iff2[THEN iffD2]) |
|
104 |
fix A assume A: "A \<in> leb" |
|
105 |
moreover have "indicator (UNIV - A) = (\<lambda>x. 1 - indicator A x :: real)" |
|
41654 | 106 |
by (auto simp: fun_eq_iff indicator_def) |
47694 | 107 |
ultimately show "UNIV - A \<in> leb" |
108 |
using A by (auto intro!: integrable_sub simp: cube_def leb_def) |
|
41654 | 109 |
next |
47694 | 110 |
fix n show "{} \<in> leb" |
111 |
by (auto simp: cube_def indicator_def[abs_def] leb_def) |
|
41654 | 112 |
next |
47694 | 113 |
fix A :: "nat \<Rightarrow> _" assume A: "range A \<subseteq> leb" |
114 |
have "\<forall>n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n" (is "\<forall>n. ?g integrable_on _") |
|
115 |
proof (intro dominated_convergence[where g="?g"] ballI allI) |
|
116 |
fix k n show "(indicator (\<Union>i<k. A i) :: _ \<Rightarrow> real) integrable_on cube n" |
|
41654 | 117 |
proof (induct k) |
118 |
case (Suc k) |
|
119 |
have *: "(\<Union> i<Suc k. A i) = (\<Union> i<k. A i) \<union> A k" |
|
120 |
unfolding lessThan_Suc UN_insert by auto |
|
121 |
have *: "(\<lambda>x. max (indicator (\<Union> i<k. A i) x) (indicator (A k) x) :: real) = |
|
122 |
indicator (\<Union> i<Suc k. A i)" (is "(\<lambda>x. max (?f x) (?g x)) = _") |
|
123 |
by (auto simp: fun_eq_iff * indicator_def) |
|
124 |
show ?case |
|
47694 | 125 |
using absolutely_integrable_max[of ?f "cube n" ?g] A Suc |
126 |
by (simp add: * leb_def subset_eq) |
|
41654 | 127 |
qed auto |
128 |
qed (auto intro: LIMSEQ_indicator_UN simp: cube_def) |
|
47694 | 129 |
then show "(\<Union>i. A i) \<in> leb" by (auto simp: leb_def) |
41654 | 130 |
qed simp |
38656 | 131 |
|
47694 | 132 |
lemma sets_lebesgue: "sets lebesgue = {A. \<forall>n. (indicator A :: _ \<Rightarrow> real) integrable_on cube n}" |
133 |
unfolding lebesgue_def sigma_algebra.sets_measure_of_eq[OF sigma_algebra_lebesgue] .. |
|
134 |
||
135 |
lemma lebesgueD: "A \<in> sets lebesgue \<Longrightarrow> (indicator A :: _ \<Rightarrow> real) integrable_on cube n" |
|
136 |
unfolding sets_lebesgue by simp |
|
137 |
||
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
138 |
lemma emeasure_lebesgue: |
47694 | 139 |
assumes "A \<in> sets lebesgue" |
140 |
shows "emeasure lebesgue A = (SUP n. ereal (integral (cube n) (indicator A)))" |
|
141 |
(is "_ = ?\<mu> A") |
|
142 |
proof (rule emeasure_measure_of[OF lebesgue_def]) |
|
41654 | 143 |
have *: "indicator {} = (\<lambda>x. 0 :: real)" by (simp add: fun_eq_iff) |
47694 | 144 |
show "positive (sets lebesgue) ?\<mu>" |
145 |
proof (unfold positive_def, intro conjI ballI) |
|
146 |
show "?\<mu> {} = 0" by (simp add: integral_0 *) |
|
147 |
fix A :: "'a set" assume "A \<in> sets lebesgue" then show "0 \<le> ?\<mu> A" |
|
148 |
by (auto intro!: SUP_upper2 Integration.integral_nonneg simp: sets_lebesgue) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
149 |
qed |
40859 | 150 |
next |
47694 | 151 |
show "countably_additive (sets lebesgue) ?\<mu>" |
41654 | 152 |
proof (intro countably_additive_def[THEN iffD2] allI impI) |
47694 | 153 |
fix A :: "nat \<Rightarrow> 'a set" assume rA: "range A \<subseteq> sets lebesgue" "disjoint_family A" |
41654 | 154 |
then have A[simp, intro]: "\<And>i n. (indicator (A i) :: _ \<Rightarrow> real) integrable_on cube n" |
155 |
by (auto dest: lebesgueD) |
|
46731 | 156 |
let ?m = "\<lambda>n i. integral (cube n) (indicator (A i) :: _\<Rightarrow>real)" |
157 |
let ?M = "\<lambda>n I. integral (cube n) (indicator (\<Union>i\<in>I. A i) :: _\<Rightarrow>real)" |
|
47694 | 158 |
have nn[simp, intro]: "\<And>i n. 0 \<le> ?m n i" by (auto intro!: Integration.integral_nonneg) |
41654 | 159 |
assume "(\<Union>i. A i) \<in> sets lebesgue" |
160 |
then have UN_A[simp, intro]: "\<And>i n. (indicator (\<Union>i. A i) :: _ \<Rightarrow> real) integrable_on cube n" |
|
47694 | 161 |
by (auto simp: sets_lebesgue) |
162 |
show "(\<Sum>n. ?\<mu> (A n)) = ?\<mu> (\<Union>i. A i)" |
|
49777 | 163 |
proof (subst suminf_SUP_eq, safe intro!: incseq_SucI) |
43920 | 164 |
fix i n show "ereal (?m n i) \<le> ereal (?m (Suc n) i)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
165 |
using cube_subset[of n "Suc n"] by (auto intro!: integral_subset_le incseq_SucI) |
41654 | 166 |
next |
43920 | 167 |
fix i n show "0 \<le> ereal (?m n i)" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
168 |
using rA unfolding lebesgue_def |
44928
7ef6505bde7f
renamed Complete_Lattices lemmas, removed legacy names
hoelzl
parents:
44890
diff
changeset
|
169 |
by (auto intro!: SUP_upper2 integral_nonneg) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
170 |
next |
43920 | 171 |
show "(SUP n. \<Sum>i. ereal (?m n i)) = (SUP n. ereal (?M n UNIV))" |
56218
1c3f1f2431f9
elongated INFI and SUPR, to reduced risk of confusing theorems names in the future while still being consistent with INTER and UNION
haftmann
parents:
56193
diff
changeset
|
172 |
proof (intro arg_cong[where f="SUPREMUM UNIV"] ext sums_unique[symmetric] sums_ereal[THEN iffD2] sums_def[THEN iffD2]) |
41654 | 173 |
fix n |
174 |
have "\<And>m. (UNION {..<m} A) \<in> sets lebesgue" using rA by auto |
|
175 |
from lebesgueD[OF this] |
|
176 |
have "(\<lambda>m. ?M n {..< m}) ----> ?M n UNIV" |
|
177 |
(is "(\<lambda>m. integral _ (?A m)) ----> ?I") |
|
178 |
by (intro dominated_convergence(2)[where f="?A" and h="\<lambda>x. 1::real"]) |
|
179 |
(auto intro: LIMSEQ_indicator_UN simp: cube_def) |
|
180 |
moreover |
|
181 |
{ fix m have *: "(\<Sum>x<m. ?m n x) = ?M n {..< m}" |
|
182 |
proof (induct m) |
|
183 |
case (Suc m) |
|
184 |
have "(\<Union>i<m. A i) \<in> sets lebesgue" using rA by auto |
|
185 |
then have "(indicator (\<Union>i<m. A i) :: _\<Rightarrow>real) integrable_on (cube n)" |
|
186 |
by (auto dest!: lebesgueD) |
|
187 |
moreover |
|
188 |
have "(\<Union>i<m. A i) \<inter> A m = {}" |
|
189 |
using rA(2)[unfolded disjoint_family_on_def, THEN bspec, of m] |
|
190 |
by auto |
|
191 |
then have "\<And>x. indicator (\<Union>i<Suc m. A i) x = |
|
192 |
indicator (\<Union>i<m. A i) x + (indicator (A m) x :: real)" |
|
193 |
by (auto simp: indicator_add lessThan_Suc ac_simps) |
|
194 |
ultimately show ?case |
|
47694 | 195 |
using Suc A by (simp add: Integration.integral_add[symmetric]) |
41654 | 196 |
qed auto } |
56193
c726ecfb22b6
cleanup Series: sorted according to typeclass hierarchy, use {..<_} instead of {0..<_}
hoelzl
parents:
56188
diff
changeset
|
197 |
ultimately show "(\<lambda>m. \<Sum>x<m. ?m n x) ----> ?M n UNIV" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
198 |
by (simp add: atLeast0LessThan) |
41654 | 199 |
qed |
200 |
qed |
|
201 |
qed |
|
47694 | 202 |
qed (auto, fact) |
40859 | 203 |
|
41654 | 204 |
lemma lebesgueI_borel[intro, simp]: |
205 |
fixes s::"'a::ordered_euclidean_space set" |
|
40859 | 206 |
assumes "s \<in> sets borel" shows "s \<in> sets lebesgue" |
41654 | 207 |
proof - |
47694 | 208 |
have "s \<in> sigma_sets (space lebesgue) (range (\<lambda>(a, b). {a .. (b :: 'a\<Colon>ordered_euclidean_space)}))" |
209 |
using assms by (simp add: borel_eq_atLeastAtMost) |
|
210 |
also have "\<dots> \<subseteq> sets lebesgue" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
211 |
proof (safe intro!: sets.sigma_sets_subset lebesgueI) |
41654 | 212 |
fix n :: nat and a b :: 'a |
213 |
show "(indicator {a..b} :: 'a\<Rightarrow>real) integrable_on cube n" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
214 |
unfolding integrable_on_def using has_integral_interval_cube[of a b] by auto |
41654 | 215 |
qed |
47694 | 216 |
finally show ?thesis . |
38656 | 217 |
qed |
218 |
||
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
219 |
lemma borel_measurable_lebesgueI: |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
220 |
"f \<in> borel_measurable borel \<Longrightarrow> f \<in> borel_measurable lebesgue" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
221 |
unfolding measurable_def by simp |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
222 |
|
40859 | 223 |
lemma lebesgueI_negligible[dest]: fixes s::"'a::ordered_euclidean_space set" |
224 |
assumes "negligible s" shows "s \<in> sets lebesgue" |
|
56188 | 225 |
using assms by (force simp: cbox_interval[symmetric] cube_def integrable_on_def negligible_def intro!: lebesgueI) |
38656 | 226 |
|
41654 | 227 |
lemma lmeasure_eq_0: |
47694 | 228 |
fixes S :: "'a::ordered_euclidean_space set" |
229 |
assumes "negligible S" shows "emeasure lebesgue S = 0" |
|
40859 | 230 |
proof - |
41654 | 231 |
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
|
232 |
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
|
233 |
by (intro integral_unique some1_equality ex_ex1I) |
56188 | 234 |
(auto simp: cube_def negligible_def cbox_interval[symmetric]) |
47694 | 235 |
then show ?thesis |
236 |
using assms by (simp add: emeasure_lebesgue lebesgueI_negligible) |
|
40859 | 237 |
qed |
238 |
||
239 |
lemma lmeasure_iff_LIMSEQ: |
|
47694 | 240 |
assumes A: "A \<in> sets lebesgue" and "0 \<le> m" |
241 |
shows "emeasure lebesgue A = ereal m \<longleftrightarrow> (\<lambda>n. integral (cube n) (indicator A :: _ \<Rightarrow> real)) ----> m" |
|
242 |
proof (subst emeasure_lebesgue[OF A], intro SUP_eq_LIMSEQ) |
|
41654 | 243 |
show "mono (\<lambda>n. integral (cube n) (indicator A::_=>real))" |
244 |
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
|
245 |
qed |
38656 | 246 |
|
41654 | 247 |
lemma lmeasure_finite_has_integral: |
248 |
fixes s :: "'a::ordered_euclidean_space set" |
|
49777 | 249 |
assumes "s \<in> sets lebesgue" "emeasure lebesgue s = ereal m" |
41654 | 250 |
shows "(indicator s has_integral m) UNIV" |
251 |
proof - |
|
252 |
let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real" |
|
49777 | 253 |
have "0 \<le> m" |
254 |
using emeasure_nonneg[of lebesgue s] `emeasure lebesgue s = ereal m` by simp |
|
41654 | 255 |
have **: "(?I s) integrable_on UNIV \<and> (\<lambda>k. integral UNIV (?I (s \<inter> cube k))) ----> integral UNIV (?I s)" |
256 |
proof (intro monotone_convergence_increasing allI ballI) |
|
257 |
have LIMSEQ: "(\<lambda>n. integral (cube n) (?I s)) ----> m" |
|
49777 | 258 |
using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1) `0 \<le> m`] . |
41654 | 259 |
{ fix n have "integral (cube n) (?I s) \<le> m" |
260 |
using cube_subset assms |
|
261 |
by (intro incseq_le[where L=m] LIMSEQ incseq_def[THEN iffD2] integral_subset_le allI impI) |
|
262 |
(auto dest!: lebesgueD) } |
|
263 |
moreover |
|
264 |
{ fix n have "0 \<le> integral (cube n) (?I s)" |
|
47694 | 265 |
using assms by (auto dest!: lebesgueD intro!: Integration.integral_nonneg) } |
41654 | 266 |
ultimately |
267 |
show "bounded {integral UNIV (?I (s \<inter> cube k)) |k. True}" |
|
268 |
unfolding bounded_def |
|
269 |
apply (rule_tac exI[of _ 0]) |
|
270 |
apply (rule_tac exI[of _ m]) |
|
271 |
by (auto simp: dist_real_def integral_indicator_UNIV) |
|
272 |
fix k show "?I (s \<inter> cube k) integrable_on UNIV" |
|
273 |
unfolding integrable_indicator_UNIV using assms by (auto dest!: lebesgueD) |
|
274 |
fix x show "?I (s \<inter> cube k) x \<le> ?I (s \<inter> cube (Suc k)) x" |
|
275 |
using cube_subset[of k "Suc k"] by (auto simp: indicator_def) |
|
276 |
next |
|
277 |
fix x :: 'a |
|
278 |
from mem_big_cube obtain k where k: "x \<in> cube k" . |
|
279 |
{ fix n have "?I (s \<inter> cube (n + k)) x = ?I s x" |
|
280 |
using k cube_subset[of k "n + k"] by (auto simp: indicator_def) } |
|
281 |
note * = this |
|
282 |
show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x" |
|
283 |
by (rule LIMSEQ_offset[where k=k]) (auto simp: *) |
|
284 |
qed |
|
285 |
note ** = conjunctD2[OF this] |
|
286 |
have m: "m = integral UNIV (?I s)" |
|
287 |
apply (intro LIMSEQ_unique[OF _ **(2)]) |
|
49777 | 288 |
using assms(2) unfolding lmeasure_iff_LIMSEQ[OF assms(1) `0 \<le> m`] integral_indicator_UNIV . |
41654 | 289 |
show ?thesis |
290 |
unfolding m by (intro integrable_integral **) |
|
38656 | 291 |
qed |
292 |
||
47694 | 293 |
lemma lmeasure_finite_integrable: assumes s: "s \<in> sets lebesgue" and "emeasure lebesgue s \<noteq> \<infinity>" |
41654 | 294 |
shows "(indicator s :: _ \<Rightarrow> real) integrable_on UNIV" |
47694 | 295 |
proof (cases "emeasure lebesgue s") |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
296 |
case (real m) |
47694 | 297 |
with lmeasure_finite_has_integral[OF `s \<in> sets lebesgue` this] emeasure_nonneg[of lebesgue s] |
41654 | 298 |
show ?thesis unfolding integrable_on_def by auto |
47694 | 299 |
qed (insert assms emeasure_nonneg[of lebesgue s], auto) |
38656 | 300 |
|
41654 | 301 |
lemma has_integral_lebesgue: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV" |
302 |
shows "s \<in> sets lebesgue" |
|
303 |
proof (intro lebesgueI) |
|
304 |
let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real" |
|
305 |
fix n show "(?I s) integrable_on cube n" unfolding cube_def |
|
306 |
proof (intro integrable_on_subinterval) |
|
307 |
show "(?I s) integrable_on UNIV" |
|
308 |
unfolding integrable_on_def using assms by auto |
|
309 |
qed auto |
|
38656 | 310 |
qed |
311 |
||
41654 | 312 |
lemma has_integral_lmeasure: assumes "((indicator s :: _\<Rightarrow>real) has_integral m) UNIV" |
47694 | 313 |
shows "emeasure lebesgue s = ereal m" |
41654 | 314 |
proof (intro lmeasure_iff_LIMSEQ[THEN iffD2]) |
315 |
let ?I = "indicator :: 'a set \<Rightarrow> 'a \<Rightarrow> real" |
|
316 |
show "s \<in> sets lebesgue" using has_integral_lebesgue[OF assms] . |
|
317 |
show "0 \<le> m" using assms by (rule has_integral_nonneg) auto |
|
318 |
have "(\<lambda>n. integral UNIV (?I (s \<inter> cube n))) ----> integral UNIV (?I s)" |
|
319 |
proof (intro dominated_convergence(2) ballI) |
|
320 |
show "(?I s) integrable_on UNIV" unfolding integrable_on_def using assms by auto |
|
321 |
fix n show "?I (s \<inter> cube n) integrable_on UNIV" |
|
322 |
unfolding integrable_indicator_UNIV using `s \<in> sets lebesgue` by (auto dest: lebesgueD) |
|
323 |
fix x show "norm (?I (s \<inter> cube n) x) \<le> ?I s x" by (auto simp: indicator_def) |
|
324 |
next |
|
325 |
fix x :: 'a |
|
326 |
from mem_big_cube obtain k where k: "x \<in> cube k" . |
|
327 |
{ fix n have "?I (s \<inter> cube (n + k)) x = ?I s x" |
|
328 |
using k cube_subset[of k "n + k"] by (auto simp: indicator_def) } |
|
329 |
note * = this |
|
330 |
show "(\<lambda>k. ?I (s \<inter> cube k) x) ----> ?I s x" |
|
331 |
by (rule LIMSEQ_offset[where k=k]) (auto simp: *) |
|
332 |
qed |
|
333 |
then show "(\<lambda>n. integral (cube n) (?I s)) ----> m" |
|
334 |
unfolding integral_unique[OF assms] integral_indicator_UNIV by simp |
|
335 |
qed |
|
336 |
||
337 |
lemma has_integral_iff_lmeasure: |
|
49777 | 338 |
"(indicator A has_integral m) UNIV \<longleftrightarrow> (A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m)" |
40859 | 339 |
proof |
41654 | 340 |
assume "(indicator A has_integral m) UNIV" |
341 |
with has_integral_lmeasure[OF this] has_integral_lebesgue[OF this] |
|
49777 | 342 |
show "A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m" |
41654 | 343 |
by (auto intro: has_integral_nonneg) |
40859 | 344 |
next |
49777 | 345 |
assume "A \<in> sets lebesgue \<and> emeasure lebesgue A = ereal m" |
41654 | 346 |
then show "(indicator A has_integral m) UNIV" by (intro lmeasure_finite_has_integral) auto |
38656 | 347 |
qed |
348 |
||
41654 | 349 |
lemma lmeasure_eq_integral: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" |
47694 | 350 |
shows "emeasure lebesgue s = ereal (integral UNIV (indicator s))" |
41654 | 351 |
using assms unfolding integrable_on_def |
352 |
proof safe |
|
353 |
fix y :: real assume "(indicator s has_integral y) UNIV" |
|
354 |
from this[unfolded has_integral_iff_lmeasure] integral_unique[OF this] |
|
47694 | 355 |
show "emeasure lebesgue s = ereal (integral UNIV (indicator s))" by simp |
40859 | 356 |
qed |
38656 | 357 |
|
358 |
lemma lebesgue_simple_function_indicator: |
|
43920 | 359 |
fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
360 |
assumes f:"simple_function lebesgue f" |
38656 | 361 |
shows "f = (\<lambda>x. (\<Sum>y \<in> f ` UNIV. y * indicator (f -` {y}) x))" |
47694 | 362 |
by (rule, subst simple_function_indicator_representation[OF f]) auto |
38656 | 363 |
|
41654 | 364 |
lemma integral_eq_lmeasure: |
47694 | 365 |
"(indicator s::_\<Rightarrow>real) integrable_on UNIV \<Longrightarrow> integral UNIV (indicator s) = real (emeasure lebesgue s)" |
41654 | 366 |
by (subst lmeasure_eq_integral) (auto intro!: integral_nonneg) |
38656 | 367 |
|
47694 | 368 |
lemma lmeasure_finite: assumes "(indicator s::_\<Rightarrow>real) integrable_on UNIV" shows "emeasure lebesgue s \<noteq> \<infinity>" |
41654 | 369 |
using lmeasure_eq_integral[OF assms] by auto |
38656 | 370 |
|
40859 | 371 |
lemma negligible_iff_lebesgue_null_sets: |
47694 | 372 |
"negligible A \<longleftrightarrow> A \<in> null_sets lebesgue" |
40859 | 373 |
proof |
374 |
assume "negligible A" |
|
375 |
from this[THEN lebesgueI_negligible] this[THEN lmeasure_eq_0] |
|
47694 | 376 |
show "A \<in> null_sets lebesgue" by auto |
40859 | 377 |
next |
47694 | 378 |
assume A: "A \<in> null_sets lebesgue" |
379 |
then have *:"((indicator A) has_integral (0::real)) UNIV" using lmeasure_finite_has_integral[of A] |
|
380 |
by (auto simp: null_sets_def) |
|
41654 | 381 |
show "negligible A" unfolding negligible_def |
382 |
proof (intro allI) |
|
383 |
fix a b :: 'a |
|
56188 | 384 |
have integrable: "(indicator A :: _\<Rightarrow>real) integrable_on cbox a b" |
385 |
by (intro integrable_on_subcbox has_integral_integrable) (auto intro: *) |
|
386 |
then have "integral (cbox a b) (indicator A) \<le> (integral UNIV (indicator A) :: real)" |
|
47694 | 387 |
using * by (auto intro!: integral_subset_le) |
56188 | 388 |
moreover have "(0::real) \<le> integral (cbox a b) (indicator A)" |
41654 | 389 |
using integrable by (auto intro!: integral_nonneg) |
56188 | 390 |
ultimately have "integral (cbox a b) (indicator A) = (0::real)" |
41654 | 391 |
using integral_unique[OF *] by auto |
56188 | 392 |
then show "(indicator A has_integral (0::real)) (cbox a b)" |
41654 | 393 |
using integrable_integral[OF integrable] by simp |
394 |
qed |
|
395 |
qed |
|
396 |
||
47694 | 397 |
lemma lmeasure_UNIV[intro]: "emeasure lebesgue (UNIV::'a::ordered_euclidean_space set) = \<infinity>" |
398 |
proof (simp add: emeasure_lebesgue, intro SUP_PInfty bexI) |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
399 |
fix n :: nat |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
400 |
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
|
401 |
moreover |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
402 |
{ 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
|
403 |
proof (cases n) |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
404 |
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
|
405 |
next |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
406 |
case (Suc n') |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
407 |
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
|
408 |
also have "(2 * real n)^1 \<le> (2 * real n) ^ DIM('a)" |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
409 |
using Suc DIM_positive[where 'a='a] |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
410 |
by (intro power_increasing) (auto simp: real_of_nat_Suc simp del: DIM_positive) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
411 |
finally show ?thesis . |
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
412 |
qed } |
43920 | 413 |
ultimately show "ereal (real n) \<le> ereal (integral (cube n) (indicator UNIV::'a\<Rightarrow>real))" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
414 |
using integral_const DIM_positive[where 'a='a] |
56188 | 415 |
by (auto simp: cube_def content_cbox_cases setprod_constant setsum_negf cbox_interval[symmetric]) |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
416 |
qed simp |
40859 | 417 |
|
49777 | 418 |
lemma lmeasure_complete: "A \<subseteq> B \<Longrightarrow> B \<in> null_sets lebesgue \<Longrightarrow> A \<in> null_sets lebesgue" |
419 |
unfolding negligible_iff_lebesgue_null_sets[symmetric] by (auto simp: negligible_subset) |
|
420 |
||
40859 | 421 |
lemma |
422 |
fixes a b ::"'a::ordered_euclidean_space" |
|
47694 | 423 |
shows lmeasure_atLeastAtMost[simp]: "emeasure lebesgue {a..b} = ereal (content {a..b})" |
41654 | 424 |
proof - |
425 |
have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV" |
|
56188 | 426 |
unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def] cbox_interval[symmetric]) |
41654 | 427 |
from lmeasure_eq_integral[OF this] show ?thesis unfolding integral_indicator_UNIV |
56188 | 428 |
by (simp add: indicator_def [abs_def] cbox_interval[symmetric]) |
40859 | 429 |
qed |
430 |
||
57138
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
431 |
lemma |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
432 |
fixes a b ::"'a::ordered_euclidean_space" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
433 |
shows lmeasure_cbox[simp]: "emeasure lebesgue (cbox a b) = ereal (content (cbox a b))" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
434 |
proof - |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
435 |
have "(indicator (UNIV \<inter> {a..b})::_\<Rightarrow>real) integrable_on UNIV" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
436 |
unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def] cbox_interval[symmetric]) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
437 |
from lmeasure_eq_integral[OF this] show ?thesis unfolding integral_indicator_UNIV |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
438 |
by (simp add: indicator_def [abs_def] cbox_interval[symmetric]) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
439 |
qed |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
440 |
|
40859 | 441 |
lemma lmeasure_singleton[simp]: |
47694 | 442 |
fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0" |
41654 | 443 |
using lmeasure_atLeastAtMost[of a a] by simp |
40859 | 444 |
|
49777 | 445 |
lemma AE_lebesgue_singleton: |
446 |
fixes a :: "'a::ordered_euclidean_space" shows "AE x in lebesgue. x \<noteq> a" |
|
447 |
by (rule AE_I[where N="{a}"]) auto |
|
448 |
||
40859 | 449 |
declare content_real[simp] |
450 |
||
451 |
lemma |
|
452 |
fixes a b :: real |
|
453 |
shows lmeasure_real_greaterThanAtMost[simp]: |
|
47694 | 454 |
"emeasure lebesgue {a <.. b} = ereal (if a \<le> b then b - a else 0)" |
49777 | 455 |
proof - |
456 |
have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b}" |
|
457 |
using AE_lebesgue_singleton[of a] |
|
458 |
by (intro emeasure_eq_AE) auto |
|
40859 | 459 |
then show ?thesis by auto |
49777 | 460 |
qed |
40859 | 461 |
|
462 |
lemma |
|
463 |
fixes a b :: real |
|
464 |
shows lmeasure_real_atLeastLessThan[simp]: |
|
47694 | 465 |
"emeasure lebesgue {a ..< b} = ereal (if a \<le> b then b - a else 0)" |
49777 | 466 |
proof - |
467 |
have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b}" |
|
468 |
using AE_lebesgue_singleton[of b] |
|
469 |
by (intro emeasure_eq_AE) auto |
|
41654 | 470 |
then show ?thesis by auto |
49777 | 471 |
qed |
41654 | 472 |
|
473 |
lemma |
|
474 |
fixes a b :: real |
|
475 |
shows lmeasure_real_greaterThanLessThan[simp]: |
|
47694 | 476 |
"emeasure lebesgue {a <..< b} = ereal (if a \<le> b then b - a else 0)" |
49777 | 477 |
proof - |
478 |
have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a .. b}" |
|
479 |
using AE_lebesgue_singleton[of a] AE_lebesgue_singleton[of b] |
|
480 |
by (intro emeasure_eq_AE) auto |
|
40859 | 481 |
then show ?thesis by auto |
49777 | 482 |
qed |
40859 | 483 |
|
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41704
diff
changeset
|
484 |
subsection {* Lebesgue-Borel measure *} |
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41704
diff
changeset
|
485 |
|
47694 | 486 |
definition "lborel = measure_of UNIV (sets borel) (emeasure lebesgue)" |
41689
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
hoelzl
parents:
41661
diff
changeset
|
487 |
|
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
|
488 |
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
|
489 |
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
|
490 |
and sets_lborel[simp]: "sets lborel = sets borel" |
47694 | 491 |
and measurable_lborel1[simp]: "measurable lborel = measurable borel" |
492 |
and measurable_lborel2[simp]: "measurable A lborel = measurable A borel" |
|
50244
de72bbe42190
qualified interpretation of sigma_algebra, to avoid name clashes
immler
parents:
50123
diff
changeset
|
493 |
using sets.sigma_sets_eq[of borel] |
47694 | 494 |
by (auto simp add: lborel_def measurable_def[abs_def]) |
40859 | 495 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
496 |
(* TODO: switch these rules! *) |
47694 | 497 |
lemma emeasure_lborel[simp]: "A \<in> sets borel \<Longrightarrow> emeasure lborel A = emeasure lebesgue A" |
498 |
by (rule emeasure_measure_of[OF lborel_def]) |
|
499 |
(auto simp: positive_def emeasure_nonneg countably_additive_def intro!: suminf_emeasure) |
|
40859 | 500 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
501 |
lemma measure_lborel[simp]: "A \<in> sets borel \<Longrightarrow> measure lborel A = measure lebesgue A" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
502 |
unfolding measure_def by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
503 |
|
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
|
504 |
interpretation lborel: sigma_finite_measure lborel |
47694 | 505 |
proof (default, intro conjI exI[of _ "\<lambda>n. cube n"]) |
506 |
show "range cube \<subseteq> sets lborel" by (auto intro: borel_closed) |
|
507 |
{ fix x :: 'a have "\<exists>n. x\<in>cube n" using mem_big_cube by auto } |
|
508 |
then show "(\<Union>i. cube i) = (space lborel :: 'a set)" using mem_big_cube by auto |
|
509 |
show "\<forall>i. emeasure lborel (cube i) \<noteq> \<infinity>" by (simp add: cube_def) |
|
510 |
qed |
|
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
|
511 |
|
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
|
512 |
interpretation lebesgue: sigma_finite_measure lebesgue |
40859 | 513 |
proof |
47694 | 514 |
from lborel.sigma_finite guess A :: "nat \<Rightarrow> 'a set" .. |
515 |
then show "\<exists>A::nat \<Rightarrow> 'a set. range A \<subseteq> sets lebesgue \<and> (\<Union>i. A i) = space lebesgue \<and> (\<forall>i. emeasure lebesgue (A i) \<noteq> \<infinity>)" |
|
516 |
by (intro exI[of _ A]) (auto simp: subset_eq) |
|
40859 | 517 |
qed |
518 |
||
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
519 |
interpretation lborel_pair: pair_sigma_finite lborel lborel .. |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
520 |
|
49777 | 521 |
lemma Int_stable_atLeastAtMost: |
522 |
fixes x::"'a::ordered_euclidean_space" |
|
523 |
shows "Int_stable (range (\<lambda>(a, b::'a). {a..b}))" |
|
56188 | 524 |
by (auto simp: inter_interval Int_stable_def cbox_interval[symmetric]) |
49777 | 525 |
|
526 |
lemma lborel_eqI: |
|
527 |
fixes M :: "'a::ordered_euclidean_space measure" |
|
528 |
assumes emeasure_eq: "\<And>a b. emeasure M {a .. b} = content {a .. b}" |
|
529 |
assumes sets_eq: "sets M = sets borel" |
|
530 |
shows "lborel = M" |
|
531 |
proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost]) |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
532 |
let ?P = "\<Pi>\<^sub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel" |
49777 | 533 |
let ?E = "range (\<lambda>(a, b). {a..b} :: 'a set)" |
534 |
show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E" |
|
535 |
by (simp_all add: borel_eq_atLeastAtMost sets_eq) |
|
536 |
||
537 |
show "range cube \<subseteq> ?E" unfolding cube_def [abs_def] by auto |
|
538 |
{ fix x :: 'a have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastforce } |
|
539 |
then show "(\<Union>i. cube i :: 'a set) = UNIV" by auto |
|
540 |
||
541 |
{ fix i show "emeasure lborel (cube i) \<noteq> \<infinity>" unfolding cube_def by auto } |
|
542 |
{ fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X" |
|
543 |
by (auto simp: emeasure_eq) } |
|
544 |
qed |
|
545 |
||
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
546 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
547 |
(* GENEREALIZE to euclidean_spaces *) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
548 |
lemma lborel_real_affine: |
49777 | 549 |
fixes c :: real assumes "c \<noteq> 0" |
550 |
shows "lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. \<bar>c\<bar>)" (is "_ = ?D") |
|
551 |
proof (rule lborel_eqI) |
|
552 |
fix a b show "emeasure ?D {a..b} = content {a .. b}" |
|
553 |
proof cases |
|
554 |
assume "0 < c" |
|
555 |
then have "(\<lambda>x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}" |
|
556 |
by (auto simp: field_simps) |
|
557 |
with `0 < c` show ?thesis |
|
558 |
by (cases "a \<le> b") |
|
56996 | 559 |
(auto simp: field_simps emeasure_density nn_integral_distr nn_integral_cmult |
49777 | 560 |
borel_measurable_indicator' emeasure_distr) |
561 |
next |
|
562 |
assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto |
|
563 |
then have *: "(\<lambda>x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}" |
|
564 |
by (auto simp: field_simps) |
|
565 |
with `c < 0` show ?thesis |
|
566 |
by (cases "a \<le> b") |
|
56996 | 567 |
(auto simp: field_simps emeasure_density nn_integral_distr |
568 |
nn_integral_cmult borel_measurable_indicator' emeasure_distr) |
|
49777 | 569 |
qed |
570 |
qed simp |
|
571 |
||
56996 | 572 |
lemma nn_integral_real_affine: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
573 |
fixes c :: real assumes [measurable]: "f \<in> borel_measurable borel" and c: "c \<noteq> 0" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
574 |
shows "(\<integral>\<^sup>+x. f x \<partial>lborel) = \<bar>c\<bar> * (\<integral>\<^sup>+x. f (t + c * x) \<partial>lborel)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
575 |
by (subst lborel_real_affine[OF c, of t]) |
56996 | 576 |
(simp add: nn_integral_density nn_integral_distr nn_integral_cmult) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
577 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
578 |
lemma lborel_integrable_real_affine: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
579 |
fixes f :: "real \<Rightarrow> _ :: {banach, second_countable_topology}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
580 |
assumes f: "integrable lborel f" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
581 |
shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x))" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
582 |
using f f[THEN borel_measurable_integrable] unfolding integrable_iff_bounded |
56996 | 583 |
by (subst (asm) nn_integral_real_affine[where c=c and t=t]) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
584 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
585 |
lemma lborel_integrable_real_affine_iff: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
586 |
fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
587 |
shows "c \<noteq> 0 \<Longrightarrow> integrable lborel (\<lambda>x. f (t + c * x)) \<longleftrightarrow> integrable lborel f" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
588 |
using |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
589 |
lborel_integrable_real_affine[of f c t] |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
590 |
lborel_integrable_real_affine[of "\<lambda>x. f (t + c * x)" "1/c" "-t/c"] |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
591 |
by (auto simp add: field_simps) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
592 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
593 |
lemma lborel_integral_real_affine: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
594 |
fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}" and c :: real |
57166
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57138
diff
changeset
|
595 |
assumes c: "c \<noteq> 0" shows "(\<integral>x. f x \<partial> lborel) = \<bar>c\<bar> *\<^sub>R (\<integral>x. f (t + c * x) \<partial>lborel)" |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57138
diff
changeset
|
596 |
proof cases |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57138
diff
changeset
|
597 |
assume f[measurable]: "integrable lborel f" then show ?thesis |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57138
diff
changeset
|
598 |
using c f f[THEN borel_measurable_integrable] f[THEN lborel_integrable_real_affine, of c t] |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57138
diff
changeset
|
599 |
by (subst lborel_real_affine[OF c, of t]) (simp add: integral_density integral_distr) |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57138
diff
changeset
|
600 |
next |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57138
diff
changeset
|
601 |
assume "\<not> integrable lborel f" with c show ?thesis |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57138
diff
changeset
|
602 |
by (simp add: lborel_integrable_real_affine_iff not_integrable_integral_eq) |
5cfcc616d485
use 0 as integral-value for non-integrable functions, simplify a couple of rewrite rules
hoelzl
parents:
57138
diff
changeset
|
603 |
qed |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
604 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
605 |
lemma divideR_right: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
606 |
fixes x y :: "'a::real_normed_vector" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
607 |
shows "r \<noteq> 0 \<Longrightarrow> y = x /\<^sub>R r \<longleftrightarrow> r *\<^sub>R y = x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
608 |
using scaleR_cancel_left[of r y "x /\<^sub>R r"] by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
609 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
610 |
lemma integrable_on_cmult_iff2: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
611 |
fixes c :: real |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
612 |
shows "(\<lambda>x. c * f x) integrable_on s \<longleftrightarrow> c = 0 \<or> f integrable_on s" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
613 |
using integrable_cmul[of "\<lambda>x. c * f x" s "1 / c"] integrable_cmul[of f s c] |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
614 |
by (cases "c = 0") auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
615 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
616 |
lemma lborel_has_bochner_integral_real_affine_iff: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
617 |
fixes x :: "'a :: {banach, second_countable_topology}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
618 |
shows "c \<noteq> 0 \<Longrightarrow> |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
619 |
has_bochner_integral lborel f x \<longleftrightarrow> |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
620 |
has_bochner_integral lborel (\<lambda>x. f (t + c * x)) (x /\<^sub>R \<bar>c\<bar>)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
621 |
unfolding has_bochner_integral_iff lborel_integrable_real_affine_iff |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
622 |
by (simp_all add: lborel_integral_real_affine[symmetric] divideR_right cong: conj_cong) |
49777 | 623 |
|
41706
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41704
diff
changeset
|
624 |
subsection {* Lebesgue integrable implies Gauge integrable *} |
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
hoelzl
parents:
41704
diff
changeset
|
625 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
626 |
lemma has_integral_scaleR_left: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
627 |
"(f has_integral y) s \<Longrightarrow> ((\<lambda>x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) s" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
628 |
using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
629 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
630 |
lemma has_integral_mult_left: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
631 |
fixes c :: "_ :: {real_normed_algebra}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
632 |
shows "(f has_integral y) s \<Longrightarrow> ((\<lambda>x. f x * c) has_integral (y * c)) s" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
633 |
using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
634 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
635 |
(* GENERALIZE Integration.dominated_convergence, then generalize the following theorems *) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
636 |
lemma has_integral_dominated_convergence: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
637 |
fixes f :: "nat \<Rightarrow> 'n::ordered_euclidean_space \<Rightarrow> real" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
638 |
assumes "\<And>k. (f k has_integral y k) s" "h integrable_on s" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
639 |
"\<And>k. \<forall>x\<in>s. norm (f k x) \<le> h x" "\<forall>x\<in>s. (\<lambda>k. f k x) ----> g x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
640 |
and x: "y ----> x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
641 |
shows "(g has_integral x) s" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
642 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
643 |
have int_f: "\<And>k. (f k) integrable_on s" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
644 |
using assms by (auto simp: integrable_on_def) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
645 |
have "(g has_integral (integral s g)) s" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
646 |
by (intro integrable_integral dominated_convergence[OF int_f assms(2)]) fact+ |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
647 |
moreover have "integral s g = x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
648 |
proof (rule LIMSEQ_unique) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
649 |
show "(\<lambda>i. integral s (f i)) ----> x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
650 |
using integral_unique[OF assms(1)] x by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
651 |
show "(\<lambda>i. integral s (f i)) ----> integral s g" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
652 |
by (intro dominated_convergence[OF int_f assms(2)]) fact+ |
41654 | 653 |
qed |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
654 |
ultimately show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
655 |
by simp |
40859 | 656 |
qed |
657 |
||
56996 | 658 |
lemma nn_integral_has_integral: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
659 |
fixes f::"'a::ordered_euclidean_space \<Rightarrow> real" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
660 |
assumes f: "f \<in> borel_measurable lebesgue" "\<And>x. 0 \<le> f x" "(\<integral>\<^sup>+x. f x \<partial>lebesgue) = ereal r" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
661 |
shows "(f has_integral r) UNIV" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
662 |
using f proof (induct arbitrary: r rule: borel_measurable_induct_real) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
663 |
case (set A) then show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
664 |
by (auto simp add: ereal_indicator has_integral_iff_lmeasure) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
665 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
666 |
case (mult g c) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
667 |
then have "ereal c * (\<integral>\<^sup>+ x. g x \<partial>lebesgue) = ereal r" |
56996 | 668 |
by (subst nn_integral_cmult[symmetric]) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
669 |
then obtain r' where "(c = 0 \<and> r = 0) \<or> ((\<integral>\<^sup>+ x. ereal (g x) \<partial>lebesgue) = ereal r' \<and> r = c * r')" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
670 |
by (cases "\<integral>\<^sup>+ x. ereal (g x) \<partial>lebesgue") (auto split: split_if_asm) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
671 |
with mult show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
672 |
by (auto intro!: has_integral_cmult_real) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
673 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
674 |
case (add g h) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
675 |
moreover |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
676 |
then have "(\<integral>\<^sup>+ x. h x + g x \<partial>lebesgue) = (\<integral>\<^sup>+ x. h x \<partial>lebesgue) + (\<integral>\<^sup>+ x. g x \<partial>lebesgue)" |
56996 | 677 |
unfolding plus_ereal.simps[symmetric] by (subst nn_integral_add) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
678 |
with add obtain a b where "(\<integral>\<^sup>+ x. h x \<partial>lebesgue) = ereal a" "(\<integral>\<^sup>+ x. g x \<partial>lebesgue) = ereal b" "r = a + b" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
679 |
by (cases "\<integral>\<^sup>+ x. h x \<partial>lebesgue" "\<integral>\<^sup>+ x. g x \<partial>lebesgue" rule: ereal2_cases) auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
680 |
ultimately show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
681 |
by (auto intro!: has_integral_add) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
682 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
683 |
case (seq U) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
684 |
note seq(1)[measurable] and f[measurable] |
40859 | 685 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
686 |
{ fix i x |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
687 |
have "U i x \<le> f x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
688 |
using seq(5) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
689 |
apply (rule LIMSEQ_le_const) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
690 |
using seq(4) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
691 |
apply (auto intro!: exI[of _ i] simp: incseq_def le_fun_def) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
692 |
done } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
693 |
note U_le_f = this |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
694 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
695 |
{ fix i |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
696 |
have "(\<integral>\<^sup>+x. ereal (U i x) \<partial>lebesgue) \<le> (\<integral>\<^sup>+x. ereal (f x) \<partial>lebesgue)" |
56996 | 697 |
using U_le_f by (intro nn_integral_mono) simp |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
698 |
then obtain p where "(\<integral>\<^sup>+x. U i x \<partial>lebesgue) = ereal p" "p \<le> r" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
699 |
using seq(6) by (cases "\<integral>\<^sup>+x. U i x \<partial>lebesgue") auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
700 |
moreover then have "0 \<le> p" |
56996 | 701 |
by (metis ereal_less_eq(5) nn_integral_nonneg) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
702 |
moreover note seq |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
703 |
ultimately have "\<exists>p. (\<integral>\<^sup>+x. U i x \<partial>lebesgue) = ereal p \<and> 0 \<le> p \<and> p \<le> r \<and> (U i has_integral p) UNIV" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
704 |
by auto } |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
705 |
then obtain p where p: "\<And>i. (\<integral>\<^sup>+x. ereal (U i x) \<partial>lebesgue) = ereal (p i)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
706 |
and bnd: "\<And>i. p i \<le> r" "\<And>i. 0 \<le> p i" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
707 |
and U_int: "\<And>i.(U i has_integral (p i)) UNIV" by metis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
708 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
709 |
have int_eq: "\<And>i. integral UNIV (U i) = p i" using U_int by (rule integral_unique) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
710 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
711 |
have *: "f integrable_on UNIV \<and> (\<lambda>k. integral UNIV (U k)) ----> integral UNIV f" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
712 |
proof (rule monotone_convergence_increasing) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
713 |
show "\<forall>k. U k integrable_on UNIV" using U_int by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
714 |
show "\<forall>k. \<forall>x\<in>UNIV. U k x \<le> U (Suc k) x" using `incseq U` by (auto simp: incseq_def le_fun_def) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
715 |
then show "bounded {integral UNIV (U k) |k. True}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
716 |
using bnd int_eq by (auto simp: bounded_real intro!: exI[of _ r]) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
717 |
show "\<forall>x\<in>UNIV. (\<lambda>k. U k x) ----> f x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
718 |
using seq by auto |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
719 |
qed |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
720 |
moreover have "(\<lambda>i. (\<integral>\<^sup>+x. U i x \<partial>lebesgue)) ----> (\<integral>\<^sup>+x. f x \<partial>lebesgue)" |
56996 | 721 |
using seq U_le_f by (intro nn_integral_dominated_convergence[where w=f]) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
722 |
ultimately have "integral UNIV f = r" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
723 |
by (auto simp add: int_eq p seq intro: LIMSEQ_unique) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
724 |
with * show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
725 |
by (simp add: has_integral_integral) |
40859 | 726 |
qed |
727 |
||
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
728 |
lemma has_integral_integrable_lebesgue_nonneg: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
729 |
fixes f::"'a::ordered_euclidean_space \<Rightarrow> real" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
730 |
assumes f: "integrable lebesgue f" "\<And>x. 0 \<le> f x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
731 |
shows "(f has_integral integral\<^sup>L lebesgue f) UNIV" |
56996 | 732 |
proof (rule nn_integral_has_integral) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
733 |
show "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue) = ereal (integral\<^sup>L lebesgue f)" |
56996 | 734 |
using f by (intro nn_integral_eq_integral) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
735 |
qed (insert f, auto) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
736 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
737 |
lemma has_integral_lebesgue_integral_lebesgue: |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
738 |
fixes f::"'a::ordered_euclidean_space \<Rightarrow> real" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
739 |
assumes f: "integrable lebesgue f" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
740 |
shows "(f has_integral (integral\<^sup>L lebesgue f)) UNIV" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
741 |
using f proof induct |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
742 |
case (base A c) then show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
743 |
by (auto intro!: has_integral_mult_left simp: has_integral_iff_lmeasure) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
744 |
(simp add: emeasure_eq_ereal_measure) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
745 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
746 |
case (add f g) then show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
747 |
by (auto intro!: has_integral_add) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
748 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
749 |
case (lim f s) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
750 |
show ?case |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
751 |
proof (rule has_integral_dominated_convergence) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
752 |
show "\<And>i. (s i has_integral integral\<^sup>L lebesgue (s i)) UNIV" by fact |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
753 |
show "(\<lambda>x. norm (2 * f x)) integrable_on UNIV" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
754 |
using lim by (intro has_integral_integrable[OF has_integral_integrable_lebesgue_nonneg]) auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
755 |
show "\<And>k. \<forall>x\<in>UNIV. norm (s k x) \<le> norm (2 * f x)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
756 |
using lim by (auto simp add: abs_mult) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
757 |
show "\<forall>x\<in>UNIV. (\<lambda>k. s k x) ----> f x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
758 |
using lim by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
759 |
show "(\<lambda>k. integral\<^sup>L lebesgue (s k)) ----> integral\<^sup>L lebesgue f" |
57137 | 760 |
using lim by (intro integral_dominated_convergence[where w="\<lambda>x. 2 * norm (f x)"]) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
761 |
qed |
40859 | 762 |
qed |
763 |
||
56996 | 764 |
lemma lebesgue_nn_integral_eq_borel: |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
765 |
assumes f: "f \<in> borel_measurable borel" |
56996 | 766 |
shows "integral\<^sup>N lebesgue f = integral\<^sup>N lborel f" |
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
767 |
proof - |
56996 | 768 |
from f have "integral\<^sup>N lebesgue (\<lambda>x. max 0 (f x)) = integral\<^sup>N lborel (\<lambda>x. max 0 (f x))" |
769 |
by (auto intro!: nn_integral_subalgebra[symmetric]) |
|
770 |
then show ?thesis unfolding nn_integral_max_0 . |
|
41981
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
hoelzl
parents:
41831
diff
changeset
|
771 |
qed |
41546
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
772 |
|
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
773 |
lemma lebesgue_integral_eq_borel: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
774 |
fixes f :: "_ \<Rightarrow> _::{banach, second_countable_topology}" |
41546
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
775 |
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
|
776 |
shows "integrable lebesgue f \<longleftrightarrow> integrable lborel f" (is ?P) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
777 |
and "integral\<^sup>L lebesgue f = integral\<^sup>L lborel f" (is ?I) |
41546
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
778 |
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
|
779 |
have "sets lborel \<subseteq> sets lebesgue" by auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
780 |
from integral_subalgebra[of f lborel, OF _ this _ _] |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
781 |
integrable_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
|
782 |
show ?P ?I by auto |
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
783 |
qed |
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
784 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
785 |
lemma has_integral_lebesgue_integral: |
41546
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
786 |
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
|
787 |
assumes f:"integrable lborel f" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
788 |
shows "(f has_integral (integral\<^sup>L lborel f)) UNIV" |
41546
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
789 |
proof - |
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
790 |
have borel: "f \<in> borel_measurable borel" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
791 |
using f unfolding integrable_iff_bounded by auto |
41546
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
792 |
from f show ?thesis |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
793 |
using has_integral_lebesgue_integral_lebesgue[of f] |
41546
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
794 |
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
|
795 |
qed |
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
hoelzl
parents:
41097
diff
changeset
|
796 |
|
56996 | 797 |
lemma nn_integral_bound_simple_function: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
798 |
assumes bnd: "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" "\<And>x. x \<in> space M \<Longrightarrow> f x < \<infinity>" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
799 |
assumes f[measurable]: "simple_function M f" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
800 |
assumes supp: "emeasure M {x\<in>space M. f x \<noteq> 0} < \<infinity>" |
56996 | 801 |
shows "nn_integral M f < \<infinity>" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
802 |
proof cases |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
803 |
assume "space M = {}" |
56996 | 804 |
then have "nn_integral M f = (\<integral>\<^sup>+x. 0 \<partial>M)" |
805 |
by (intro nn_integral_cong) auto |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
806 |
then show ?thesis by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
807 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
808 |
assume "space M \<noteq> {}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
809 |
with simple_functionD(1)[OF f] bnd have bnd: "0 \<le> Max (f`space M) \<and> Max (f`space M) < \<infinity>" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
810 |
by (subst Max_less_iff) (auto simp: Max_ge_iff) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
811 |
|
56996 | 812 |
have "nn_integral M f \<le> (\<integral>\<^sup>+x. Max (f`space M) * indicator {x\<in>space M. f x \<noteq> 0} x \<partial>M)" |
813 |
proof (rule nn_integral_mono) |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
814 |
fix x assume "x \<in> space M" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
815 |
with f show "f x \<le> Max (f ` space M) * indicator {x \<in> space M. f x \<noteq> 0} x" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
816 |
by (auto split: split_indicator intro!: Max_ge simple_functionD) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
817 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
818 |
also have "\<dots> < \<infinity>" |
56996 | 819 |
using bnd supp by (subst nn_integral_cmult) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
820 |
finally show ?thesis . |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
821 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
822 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
823 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
824 |
lemma |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
825 |
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real" |
49777 | 826 |
assumes f_borel: "f \<in> borel_measurable lebesgue" and nonneg: "\<And>x. 0 \<le> f x" |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
827 |
assumes I: "(f has_integral I) UNIV" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
828 |
shows integrable_has_integral_lebesgue_nonneg: "integrable lebesgue f" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
829 |
and integral_has_integral_lebesgue_nonneg: "integral\<^sup>L lebesgue f = I" |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
830 |
proof - |
49777 | 831 |
from f_borel have "(\<lambda>x. ereal (f x)) \<in> borel_measurable lebesgue" by auto |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
832 |
from borel_measurable_implies_simple_function_sequence'[OF this] guess F . note F = this |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
833 |
|
56996 | 834 |
have "(\<integral>\<^sup>+ x. ereal (f x) \<partial>lebesgue) = (SUP i. integral\<^sup>N lebesgue (F i))" |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
835 |
using F |
56996 | 836 |
by (subst nn_integral_monotone_convergence_SUP[symmetric]) |
837 |
(simp_all add: nn_integral_max_0 borel_measurable_simple_function) |
|
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
838 |
also have "\<dots> \<le> ereal I" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
839 |
proof (rule SUP_least) |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
840 |
fix i :: nat |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
841 |
|
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
842 |
{ fix z |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
843 |
from F(4)[of z] have "F i z \<le> ereal (f z)" |
54863
82acc20ded73
prefer more canonical names for lemmas on min/max
haftmann
parents:
54775
diff
changeset
|
844 |
by (metis SUP_upper UNIV_I ereal_max_0 max.absorb2 nonneg) |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
845 |
with F(5)[of i z] have "real (F i z) \<le> f z" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
846 |
by (cases "F i z") simp_all } |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
847 |
note F_bound = this |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
848 |
|
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
849 |
{ fix x :: ereal assume x: "x \<noteq> 0" "x \<in> range (F i)" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
850 |
with F(3,5)[of i] have [simp]: "real x \<noteq> 0" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
851 |
by (metis image_iff order_eq_iff real_of_ereal_le_0) |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
852 |
let ?s = "(\<lambda>n z. real x * indicator (F i -` {x} \<inter> cube n) z) :: nat \<Rightarrow> 'a \<Rightarrow> real" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
853 |
have "(\<lambda>z::'a. real x * indicator (F i -` {x}) z) integrable_on UNIV" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
854 |
proof (rule dominated_convergence(1)) |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
855 |
fix n :: nat |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
856 |
have "(\<lambda>z. indicator (F i -` {x} \<inter> cube n) z :: real) integrable_on cube n" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
857 |
using x F(1)[of i] |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
858 |
by (intro lebesgueD) (auto simp: simple_function_def) |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
859 |
then have cube: "?s n integrable_on cube n" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
860 |
by (simp add: integrable_on_cmult_iff) |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
861 |
show "?s n integrable_on UNIV" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
862 |
by (rule integrable_on_superset[OF _ _ cube]) auto |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
863 |
next |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
864 |
show "f integrable_on UNIV" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
865 |
unfolding integrable_on_def using I by auto |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
866 |
next |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
867 |
fix n from F_bound show "\<forall>x\<in>UNIV. norm (?s n x) \<le> f x" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
868 |
using nonneg F(5) by (auto split: split_indicator) |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
869 |
next |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
870 |
show "\<forall>z\<in>UNIV. (\<lambda>n. ?s n z) ----> real x * indicator (F i -` {x}) z" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
871 |
proof |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
872 |
fix z :: 'a |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
873 |
from mem_big_cube[of z] guess j . |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
874 |
then have x: "eventually (\<lambda>n. ?s n z = real x * indicator (F i -` {x}) z) sequentially" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
875 |
by (auto intro!: eventually_sequentiallyI[where c=j] dest!: cube_subset split: split_indicator) |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
876 |
then show "(\<lambda>n. ?s n z) ----> real x * indicator (F i -` {x}) z" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
877 |
by (rule Lim_eventually) |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
878 |
qed |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
879 |
qed |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
880 |
then have "(indicator (F i -` {x}) :: 'a \<Rightarrow> real) integrable_on UNIV" |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
881 |
by (simp add: integrable_on_cmult_iff) } |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
882 |
note F_finite = lmeasure_finite[OF this] |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
883 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
884 |
have F_eq: "\<And>x. F i x = ereal (norm (real (F i x)))" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
885 |
using F(3,5) by (auto simp: fun_eq_iff ereal_real image_iff eq_commute) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
886 |
have F_eq2: "\<And>x. F i x = ereal (real (F i x))" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
887 |
using F(3,5) by (auto simp: fun_eq_iff ereal_real image_iff eq_commute) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
888 |
|
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
889 |
have int: "integrable lebesgue (\<lambda>x. real (F i x))" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
890 |
unfolding integrable_iff_bounded |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
891 |
proof |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
892 |
have "(\<integral>\<^sup>+x. F i x \<partial>lebesgue) < \<infinity>" |
56996 | 893 |
proof (rule nn_integral_bound_simple_function) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
894 |
fix x::'a assume "x \<in> space lebesgue" then show "0 \<le> F i x" "F i x < \<infinity>" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
895 |
using F by (auto simp: image_iff eq_commute) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
896 |
next |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
897 |
have eq: "{x \<in> space lebesgue. F i x \<noteq> 0} = (\<Union>x\<in>F i ` space lebesgue - {0}. F i -` {x} \<inter> space lebesgue)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
898 |
by auto |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
899 |
show "emeasure lebesgue {x \<in> space lebesgue. F i x \<noteq> 0} < \<infinity>" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
900 |
unfolding eq using simple_functionD[OF F(1)] |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
901 |
by (subst setsum_emeasure[symmetric]) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
902 |
(auto simp: disjoint_family_on_def setsum_Pinfty F_finite) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
903 |
qed fact |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
904 |
with F_eq show "(\<integral>\<^sup>+x. norm (real (F i x)) \<partial>lebesgue) < \<infinity>" by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
905 |
qed (insert F(1), auto intro!: borel_measurable_real_of_ereal dest: borel_measurable_simple_function) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
906 |
then have "((\<lambda>x. real (F i x)) has_integral integral\<^sup>L lebesgue (\<lambda>x. real (F i x))) UNIV" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
907 |
by (rule has_integral_lebesgue_integral_lebesgue) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
908 |
from this I have "integral\<^sup>L lebesgue (\<lambda>x. real (F i x)) \<le> I" |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
909 |
by (rule has_integral_le) (intro ballI F_bound) |
56996 | 910 |
moreover have "integral\<^sup>N lebesgue (F i) = integral\<^sup>L lebesgue (\<lambda>x. real (F i x))" |
911 |
using int F by (subst nn_integral_eq_integral[symmetric]) (auto simp: F_eq2[symmetric] real_of_ereal_pos) |
|
912 |
ultimately show "integral\<^sup>N lebesgue (F i) \<le> ereal I" |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
913 |
by simp |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
914 |
qed |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
915 |
finally show "integrable lebesgue f" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
916 |
using f_borel by (auto simp: integrable_iff_bounded nonneg) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
917 |
from has_integral_lebesgue_integral_lebesgue[OF this] I |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
918 |
show "integral\<^sup>L lebesgue f = I" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
919 |
by (metis has_integral_unique) |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
920 |
qed |
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
921 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
922 |
lemma has_integral_iff_has_bochner_integral_lebesgue_nonneg: |
49777 | 923 |
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
924 |
shows "f \<in> borel_measurable lebesgue \<Longrightarrow> (\<And>x. 0 \<le> f x) \<Longrightarrow> |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
925 |
(f has_integral I) UNIV \<longleftrightarrow> has_bochner_integral lebesgue f I" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
926 |
by (metis has_bochner_integral_iff has_integral_unique has_integral_lebesgue_integral_lebesgue |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
927 |
integrable_has_integral_lebesgue_nonneg) |
49777 | 928 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
929 |
lemma |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
930 |
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
931 |
assumes "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x" "(f has_integral I) UNIV" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
932 |
shows integrable_has_integral_nonneg: "integrable lborel f" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
933 |
and integral_has_integral_nonneg: "integral\<^sup>L lborel f = I" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
934 |
by (metis assms borel_measurable_lebesgueI integrable_has_integral_lebesgue_nonneg lebesgue_integral_eq_borel(1)) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
935 |
(metis assms borel_measurable_lebesgueI has_integral_lebesgue_integral has_integral_unique integrable_has_integral_lebesgue_nonneg lebesgue_integral_eq_borel(1)) |
49777 | 936 |
|
937 |
subsection {* Equivalence between product spaces and euclidean spaces *} |
|
938 |
||
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
939 |
definition e2p :: "'a::ordered_euclidean_space \<Rightarrow> ('a \<Rightarrow> real)" where |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
940 |
"e2p x = (\<lambda>i\<in>Basis. x \<bullet> i)" |
49777 | 941 |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
942 |
definition p2e :: "('a \<Rightarrow> real) \<Rightarrow> 'a::ordered_euclidean_space" where |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
943 |
"p2e x = (\<Sum>i\<in>Basis. x i *\<^sub>R i)" |
49777 | 944 |
|
945 |
lemma e2p_p2e[simp]: |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
946 |
"x \<in> extensional Basis \<Longrightarrow> e2p (p2e x::'a::ordered_euclidean_space) = x" |
49777 | 947 |
by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def) |
948 |
||
949 |
lemma p2e_e2p[simp]: |
|
950 |
"p2e (e2p x) = (x::'a::ordered_euclidean_space)" |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
951 |
by (auto simp: euclidean_eq_iff[where 'a='a] p2e_def e2p_def) |
49777 | 952 |
|
953 |
interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure" |
|
954 |
by default |
|
955 |
||
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
956 |
interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "Basis" |
49777 | 957 |
by default auto |
958 |
||
959 |
lemma sets_product_borel: |
|
960 |
assumes I: "finite I" |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
961 |
shows "sets (\<Pi>\<^sub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^sub>E i\<in>I. UNIV) { \<Pi>\<^sub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G") |
49777 | 962 |
proof (subst sigma_prod_algebra_sigma_eq[where S="\<lambda>_ i::nat. {..<real i}" and E="\<lambda>_. range lessThan", OF I]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
963 |
show "sigma_sets (space (Pi\<^sub>M I (\<lambda>i. lborel))) {Pi\<^sub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G" |
49777 | 964 |
by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff) |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54257
diff
changeset
|
965 |
qed (auto simp: borel_eq_lessThan eucl_lessThan reals_Archimedean2) |
49777 | 966 |
|
50003 | 967 |
lemma measurable_e2p[measurable]: |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
968 |
"e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^sub>M (i::'a)\<in>Basis. (lborel :: real measure))" |
49777 | 969 |
proof (rule measurable_sigma_sets[OF sets_product_borel]) |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
970 |
fix A :: "('a \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^sub>E (i::'a)\<in>Basis. {..<x i} |x. True} " |
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
971 |
then obtain x where "A = (\<Pi>\<^sub>E (i::'a)\<in>Basis. {..<x i})" by auto |
54775
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54257
diff
changeset
|
972 |
then have "e2p -` A = {y :: 'a. eucl_less y (\<Sum>i\<in>Basis. x i *\<^sub>R i)}" |
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
immler
parents:
54257
diff
changeset
|
973 |
using DIM_positive by (auto simp add: set_eq_iff e2p_def eucl_less_def) |
49777 | 974 |
then show "e2p -` A \<inter> space (borel::'a measure) \<in> sets borel" by simp |
975 |
qed (auto simp: e2p_def) |
|
976 |
||
50003 | 977 |
(* FIXME: conversion in measurable prover *) |
50385 | 978 |
lemma lborelD_Collect[measurable (raw)]: "{x\<in>space borel. P x} \<in> sets borel \<Longrightarrow> {x\<in>space lborel. P x} \<in> sets lborel" by simp |
979 |
lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel" by simp |
|
50003 | 980 |
|
981 |
lemma measurable_p2e[measurable]: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
982 |
"p2e \<in> measurable (\<Pi>\<^sub>M (i::'a)\<in>Basis. (lborel :: real measure)) |
49777 | 983 |
(borel :: 'a::ordered_euclidean_space measure)" |
984 |
(is "p2e \<in> measurable ?P _") |
|
985 |
proof (safe intro!: borel_measurable_iff_halfspace_le[THEN iffD2]) |
|
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
986 |
fix x and i :: 'a |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
987 |
let ?A = "{w \<in> space ?P. (p2e w :: 'a) \<bullet> i \<le> x}" |
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
988 |
assume "i \<in> Basis" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
989 |
then have "?A = (\<Pi>\<^sub>E j\<in>Basis. if i = j then {.. x} else UNIV)" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50105
diff
changeset
|
990 |
using DIM_positive by (auto simp: space_PiM p2e_def PiE_def split: split_if_asm) |
49777 | 991 |
then show "?A \<in> sets ?P" |
992 |
by auto |
|
993 |
qed |
|
994 |
||
995 |
lemma lborel_eq_lborel_space: |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
996 |
"(lborel :: 'a measure) = distr (\<Pi>\<^sub>M (i::'a::ordered_euclidean_space)\<in>Basis. lborel) borel p2e" |
49777 | 997 |
(is "?B = ?D") |
998 |
proof (rule lborel_eqI) |
|
999 |
show "sets ?D = sets borel" by simp |
|
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
1000 |
let ?P = "(\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel)" |
49777 | 1001 |
fix a b :: 'a |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
1002 |
have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^sub>E i\<in>Basis. {a \<bullet> i .. b \<bullet> i})" |
50123
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
hoelzl
parents:
50105
diff
changeset
|
1003 |
by (auto simp: eucl_le[where 'a='a] p2e_def space_PiM PiE_def Pi_iff) |
49777 | 1004 |
have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}" |
1005 |
proof cases |
|
1006 |
assume "{a..b} \<noteq> {}" |
|
1007 |
then have "a \<le> b" |
|
56188 | 1008 |
by (simp add: eucl_le[where 'a='a]) |
50526
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
hoelzl
parents:
50418
diff
changeset
|
1009 |
then have "emeasure lborel {a..b} = (\<Prod>x\<in>Basis. emeasure lborel {a \<bullet> x .. b \<bullet> x})" |
56188 | 1010 |
by (auto simp: eucl_le[where 'a='a] content_closed_interval |
49777 | 1011 |
intro!: setprod_ereal[symmetric]) |
1012 |
also have "\<dots> = emeasure ?P (p2e -` {a..b} \<inter> space ?P)" |
|
1013 |
unfolding * by (subst lborel_space.measure_times) auto |
|
1014 |
finally show ?thesis by simp |
|
1015 |
qed simp |
|
1016 |
then show "emeasure ?D {a .. b} = content {a .. b}" |
|
1017 |
by (simp add: emeasure_distr measurable_p2e) |
|
1018 |
qed |
|
1019 |
||
1020 |
lemma borel_fubini_positiv_integral: |
|
1021 |
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal" |
|
1022 |
assumes f: "f \<in> borel_measurable borel" |
|
56996 | 1023 |
shows "integral\<^sup>N lborel f = \<integral>\<^sup>+x. f (p2e x) \<partial>(\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel)" |
1024 |
by (subst lborel_eq_lborel_space) (simp add: nn_integral_distr measurable_p2e f) |
|
49777 | 1025 |
|
1026 |
lemma borel_fubini_integrable: |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1027 |
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> _::{banach, second_countable_topology}" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
1028 |
shows "integrable lborel f \<longleftrightarrow> integrable (\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel) (\<lambda>x. f (p2e x))" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1029 |
unfolding integrable_iff_bounded |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1030 |
proof (intro conj_cong[symmetric]) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1031 |
show "((\<lambda>x. f (p2e x)) \<in> borel_measurable (Pi\<^sub>M Basis (\<lambda>i. lborel))) = (f \<in> borel_measurable lborel)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1032 |
proof |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1033 |
assume "((\<lambda>x. f (p2e x)) \<in> borel_measurable (Pi\<^sub>M Basis (\<lambda>i. lborel)))" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1034 |
then have "(\<lambda>x. f (p2e (e2p x))) \<in> borel_measurable borel" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1035 |
by measurable |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1036 |
then show "f \<in> borel_measurable lborel" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1037 |
by simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1038 |
qed simp |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1039 |
qed (simp add: borel_fubini_positiv_integral) |
49777 | 1040 |
|
1041 |
lemma borel_fubini: |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1042 |
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> _::{banach, second_countable_topology}" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1043 |
shows "f \<in> borel_measurable borel \<Longrightarrow> |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1044 |
integral\<^sup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^sub>M (i::'a)\<in>Basis. lborel))" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1045 |
by (subst lborel_eq_lborel_space) (simp add: integral_distr) |
47757
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
hoelzl
parents:
47694
diff
changeset
|
1046 |
|
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1047 |
lemma integrable_on_borel_integrable: |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1048 |
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1049 |
shows "f \<in> borel_measurable borel \<Longrightarrow> (\<And>x. 0 \<le> f x) \<Longrightarrow> f integrable_on UNIV \<Longrightarrow> integrable lborel f" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1050 |
by (metis borel_measurable_lebesgueI integrable_has_integral_nonneg integrable_on_def |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1051 |
lebesgue_integral_eq_borel(1)) |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1052 |
|
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1053 |
subsection {* Fundamental Theorem of Calculus for the Lebesgue integral *} |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1054 |
|
57138
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1055 |
lemma emeasure_bounded_finite: |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1056 |
assumes "bounded A" shows "emeasure lborel A < \<infinity>" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1057 |
proof - |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1058 |
from bounded_subset_cbox[OF `bounded A`] obtain a b where "A \<subseteq> cbox a b" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1059 |
by auto |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1060 |
then have "emeasure lborel A \<le> emeasure lborel (cbox a b)" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1061 |
by (intro emeasure_mono) auto |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1062 |
then show ?thesis |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1063 |
by auto |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1064 |
qed |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1065 |
|
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1066 |
lemma emeasure_compact_finite: "compact A \<Longrightarrow> emeasure lborel A < \<infinity>" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1067 |
using emeasure_bounded_finite[of A] by (auto intro: compact_imp_bounded) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1068 |
|
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1069 |
lemma borel_integrable_compact: |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1070 |
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> 'b::{banach, second_countable_topology}" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1071 |
assumes "compact S" "continuous_on S f" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1072 |
shows "integrable lborel (\<lambda>x. indicator S x *\<^sub>R f x)" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1073 |
proof cases |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1074 |
assume "S \<noteq> {}" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1075 |
have "continuous_on S (\<lambda>x. norm (f x))" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1076 |
using assms by (intro continuous_intros) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1077 |
from continuous_attains_sup[OF `compact S` `S \<noteq> {}` this] |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1078 |
obtain M where M: "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> M" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1079 |
by auto |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1080 |
|
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1081 |
show ?thesis |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1082 |
proof (rule integrable_bound) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1083 |
show "integrable lborel (\<lambda>x. indicator S x * M)" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1084 |
using assms by (auto intro!: emeasure_compact_finite borel_compact integrable_mult_left) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1085 |
show "(\<lambda>x. indicator S x *\<^sub>R f x) \<in> borel_measurable lborel" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1086 |
using assms by (auto intro!: borel_measurable_continuous_on_indicator borel_compact) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1087 |
show "AE x in lborel. norm (indicator S x *\<^sub>R f x) \<le> norm (indicator S x * M)" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1088 |
by (auto split: split_indicator simp: abs_real_def dest!: M) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1089 |
qed |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1090 |
qed simp |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1091 |
|
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1092 |
lemma borel_integrable_atLeastAtMost: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1093 |
fixes f :: "real \<Rightarrow> real" |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1094 |
assumes f: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1095 |
shows "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is "integrable _ ?f") |
57138
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1096 |
proof - |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1097 |
have "integrable lborel (\<lambda>x. indicator {a .. b} x *\<^sub>R f x)" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1098 |
proof (rule borel_integrable_compact) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1099 |
from f show "continuous_on {a..b} f" |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1100 |
by (auto intro: continuous_at_imp_continuous_on) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1101 |
qed simp |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1102 |
then show ?thesis |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1103 |
by (auto simp: mult_commute) |
7b3146180291
generalizd measurability on restricted space; rule for integrability on compact sets
hoelzl
parents:
57137
diff
changeset
|
1104 |
qed |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1105 |
|
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1106 |
lemma has_field_derivative_subset: |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1107 |
"(f has_field_derivative y) (at x within s) \<Longrightarrow> t \<subseteq> s \<Longrightarrow> (f has_field_derivative y) (at x within t)" |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1108 |
unfolding has_field_derivative_def by (rule has_derivative_subset) |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1109 |
|
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1110 |
lemma integral_FTC_atLeastAtMost: |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1111 |
fixes a b :: real |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1112 |
assumes "a \<le> b" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1113 |
and F: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> DERIV F x :> f x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1114 |
and f: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> isCont f x" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
1115 |
shows "integral\<^sup>L lborel (\<lambda>x. f x * indicator {a .. b} x) = F b - F a" |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1116 |
proof - |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1117 |
let ?f = "\<lambda>x. f x * indicator {a .. b} x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1118 |
have "(?f has_integral (\<integral>x. ?f x \<partial>lborel)) UNIV" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1119 |
using borel_integrable_atLeastAtMost[OF f] |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1120 |
by (rule has_integral_lebesgue_integral) |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1121 |
moreover |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1122 |
have "(f has_integral F b - F a) {a .. b}" |
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1123 |
by (intro fundamental_theorem_of_calculus) |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1124 |
(auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1125 |
intro: has_field_derivative_subset[OF F] assms(1)) |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1126 |
then have "(?f has_integral F b - F a) {a .. b}" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1127 |
by (subst has_integral_eq_eq[where g=f]) auto |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1128 |
then have "(?f has_integral F b - F a) UNIV" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1129 |
by (intro has_integral_on_superset[where t=UNIV and s="{a..b}"]) auto |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
1130 |
ultimately show "integral\<^sup>L lborel ?f = F b - F a" |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1131 |
by (rule has_integral_unique) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1132 |
qed |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1133 |
|
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1134 |
text {* |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1135 |
|
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1136 |
For the positive integral we replace continuity with Borel-measurability. |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1137 |
|
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1138 |
*} |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1139 |
|
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1140 |
lemma |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1141 |
fixes f :: "real \<Rightarrow> real" |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1142 |
assumes f_borel: "f \<in> borel_measurable borel" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1143 |
assumes f: "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" and "a \<le> b" |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1144 |
shows integral_FTC_Icc_nonneg: "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" (is ?eq) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1145 |
and integrable_FTC_Icc_nonneg: "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is ?int) |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1146 |
proof - |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1147 |
have i: "(f has_integral F b - F a) {a..b}" |
56181
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1148 |
by (intro fundamental_theorem_of_calculus) |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1149 |
(auto simp: has_field_derivative_iff_has_vector_derivative[symmetric] |
2aa0b19e74f3
unify syntax for has_derivative and differentiable
hoelzl
parents:
56020
diff
changeset
|
1150 |
intro: has_field_derivative_subset[OF f(1)] `a \<le> b`) |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1151 |
have i: "((\<lambda>x. f x * indicator {a..b} x) has_integral F b - F a) {a..b}" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1152 |
by (rule has_integral_eq[OF _ i]) auto |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1153 |
have i: "((\<lambda>x. f x * indicator {a..b} x) has_integral F b - F a) UNIV" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1154 |
by (rule has_integral_on_superset[OF _ _ i]) auto |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1155 |
from i f f_borel show ?eq |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1156 |
by (intro integral_has_integral_nonneg) (auto split: split_indicator) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1157 |
from i f f_borel show ?int |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1158 |
by (intro integrable_has_integral_nonneg) (auto split: split_indicator) |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1159 |
qed |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1160 |
|
56996 | 1161 |
lemma nn_integral_FTC_atLeastAtMost: |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1162 |
assumes "f \<in> borel_measurable borel" "\<And>x. x \<in> {a..b} \<Longrightarrow> DERIV F x :> f x" "\<And>x. x \<in> {a..b} \<Longrightarrow> 0 \<le> f x" "a \<le> b" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1163 |
shows "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1164 |
proof - |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1165 |
have "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1166 |
by (rule integrable_FTC_Icc_nonneg) fact+ |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1167 |
then have "(\<integral>\<^sup>+x. f x * indicator {a .. b} x \<partial>lborel) = (\<integral>x. f x * indicator {a .. b} x \<partial>lborel)" |
56996 | 1168 |
using assms by (intro nn_integral_eq_integral) (auto simp: indicator_def) |
56993
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1169 |
also have "(\<integral>x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a" |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1170 |
by (rule integral_FTC_Icc_nonneg) fact+ |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1171 |
finally show ?thesis |
e5366291d6aa
introduce Bochner integral: generalizes Lebesgue integral from real-valued function to functions on real-normed vector spaces
hoelzl
parents:
56218
diff
changeset
|
1172 |
unfolding ereal_indicator[symmetric] by simp |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1173 |
qed |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1174 |
|
56996 | 1175 |
lemma nn_integral_FTC_atLeast: |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1176 |
fixes f :: "real \<Rightarrow> real" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1177 |
assumes f_borel: "f \<in> borel_measurable borel" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1178 |
assumes f: "\<And>x. a \<le> x \<Longrightarrow> DERIV F x :> f x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1179 |
assumes nonneg: "\<And>x. a \<le> x \<Longrightarrow> 0 \<le> f x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1180 |
assumes lim: "(F ---> T) at_top" |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
1181 |
shows "(\<integral>\<^sup>+x. ereal (f x) * indicator {a ..} x \<partial>lborel) = T - F a" |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1182 |
proof - |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1183 |
let ?f = "\<lambda>(i::nat) (x::real). ereal (f x) * indicator {a..a + real i} x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1184 |
let ?fR = "\<lambda>x. ereal (f x) * indicator {a ..} x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1185 |
have "\<And>x. (SUP i::nat. ?f i x) = ?fR x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1186 |
proof (rule SUP_Lim_ereal) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1187 |
show "\<And>x. incseq (\<lambda>i. ?f i x)" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1188 |
using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1189 |
|
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1190 |
fix x |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1191 |
from reals_Archimedean2[of "x - a"] guess n .. |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1192 |
then have "eventually (\<lambda>n. ?f n x = ?fR x) sequentially" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1193 |
by (auto intro!: eventually_sequentiallyI[where c=n] split: split_indicator) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1194 |
then show "(\<lambda>n. ?f n x) ----> ?fR x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1195 |
by (rule Lim_eventually) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1196 |
qed |
56996 | 1197 |
then have "integral\<^sup>N lborel ?fR = (\<integral>\<^sup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)" |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1198 |
by simp |
53015
a1119cf551e8
standardized symbols via "isabelle update_sub_sup", excluding src/Pure and src/Tools/WWW_Find;
wenzelm
parents:
51478
diff
changeset
|
1199 |
also have "\<dots> = (SUP i::nat. (\<integral>\<^sup>+ x. ?f i x \<partial>lborel))" |
56996 | 1200 |
proof (rule nn_integral_monotone_convergence_SUP) |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1201 |
show "incseq ?f" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1202 |
using nonneg by (auto simp: incseq_def le_fun_def split: split_indicator) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1203 |
show "\<And>i. (?f i) \<in> borel_measurable lborel" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1204 |
using f_borel by auto |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1205 |
show "\<And>i x. 0 \<le> ?f i x" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1206 |
using nonneg by (auto split: split_indicator) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1207 |
qed |
54257
5c7a3b6b05a9
generalize SUP and INF to the syntactic type classes Sup and Inf
hoelzl
parents:
53374
diff
changeset
|
1208 |
also have "\<dots> = (SUP i::nat. ereal (F (a + real i) - F a))" |
56996 | 1209 |
by (subst nn_integral_FTC_atLeastAtMost[OF f_borel f nonneg]) auto |
50418
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1210 |
also have "\<dots> = T - F a" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1211 |
proof (rule SUP_Lim_ereal) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1212 |
show "incseq (\<lambda>n. ereal (F (a + real n) - F a))" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1213 |
proof (simp add: incseq_def, safe) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1214 |
fix m n :: nat assume "m \<le> n" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1215 |
with f nonneg show "F (a + real m) \<le> F (a + real n)" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1216 |
by (intro DERIV_nonneg_imp_nondecreasing[where f=F]) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1217 |
(simp, metis add_increasing2 order_refl order_trans real_of_nat_ge_zero) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1218 |
qed |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1219 |
have "(\<lambda>x. F (a + real x)) ----> T" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1220 |
apply (rule filterlim_compose[OF lim filterlim_tendsto_add_at_top]) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1221 |
apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl]) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1222 |
apply (rule filterlim_real_sequentially) |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1223 |
done |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1224 |
then show "(\<lambda>n. ereal (F (a + real n) - F a)) ----> ereal (T - F a)" |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1225 |
unfolding lim_ereal |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1226 |
by (intro tendsto_diff) auto |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1227 |
qed |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1228 |
finally show ?thesis . |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1229 |
qed |
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
hoelzl
parents:
50385
diff
changeset
|
1230 |
|
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1231 |
subsection {* Integration by parts *} |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1232 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1233 |
lemma integral_by_parts_integrable: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1234 |
fixes f g F G::"real \<Rightarrow> real" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1235 |
assumes "a \<le> b" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1236 |
assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1237 |
assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1238 |
assumes [intro]: "!!x. DERIV F x :> f x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1239 |
assumes [intro]: "!!x. DERIV G x :> g x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1240 |
shows "integrable lborel (\<lambda>x.((F x) * (g x) + (f x) * (G x)) * indicator {a .. b} x)" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1241 |
by (auto intro!: borel_integrable_atLeastAtMost continuous_intros) (auto intro!: DERIV_isCont) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1242 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1243 |
lemma integral_by_parts: |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1244 |
fixes f g F G::"real \<Rightarrow> real" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1245 |
assumes [arith]: "a \<le> b" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1246 |
assumes cont_f[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1247 |
assumes cont_g[intro]: "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1248 |
assumes [intro]: "!!x. DERIV F x :> f x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1249 |
assumes [intro]: "!!x. DERIV G x :> g x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1250 |
shows "(\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1251 |
= F b * G b - F a * G a - \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1252 |
proof- |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1253 |
have 0: "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) = F b * G b - F a * G a" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1254 |
by (rule integral_FTC_atLeastAtMost, auto intro!: derivative_eq_intros continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1255 |
(auto intro!: DERIV_isCont) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1256 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1257 |
have "(\<integral>x. (F x * g x + f x * G x) * indicator {a .. b} x \<partial>lborel) = |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1258 |
(\<integral>x. (F x * g x) * indicator {a .. b} x \<partial>lborel) + \<integral>x. (f x * G x) * indicator {a .. b} x \<partial>lborel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1259 |
apply (subst integral_add[symmetric]) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1260 |
apply (auto intro!: borel_integrable_atLeastAtMost continuous_intros) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1261 |
by (auto intro!: DERIV_isCont integral_cong split:split_indicator) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1262 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1263 |
thus ?thesis using 0 by auto |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1264 |
qed |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1265 |
|
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1266 |
lemma integral_by_parts': |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1267 |
fixes f g F G::"real \<Rightarrow> real" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1268 |
assumes "a \<le> b" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1269 |
assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont f x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1270 |
assumes "!!x. a \<le>x \<Longrightarrow> x\<le>b \<Longrightarrow> isCont g x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1271 |
assumes "!!x. DERIV F x :> f x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1272 |
assumes "!!x. DERIV G x :> g x" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1273 |
shows "(\<integral>x. indicator {a .. b} x *\<^sub>R (F x * g x) \<partial>lborel) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1274 |
= F b * G b - F a * G a - \<integral>x. indicator {a .. b} x *\<^sub>R (f x * G x) \<partial>lborel" |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1275 |
using integral_by_parts[OF assms] by (simp add: mult_ac) |
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1276 |
|
57275
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1277 |
lemma AE_borel_affine: |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1278 |
fixes P :: "real \<Rightarrow> bool" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1279 |
shows "c \<noteq> 0 \<Longrightarrow> Measurable.pred borel P \<Longrightarrow> AE x in lborel. P x \<Longrightarrow> AE x in lborel. P (t + c * x)" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1280 |
by (subst lborel_real_affine[where t="- t / c" and c="1 / c"]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1281 |
(simp_all add: AE_density AE_distr_iff field_simps) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1282 |
|
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1283 |
lemma has_bochner_integral_even_function: |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1284 |
fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1285 |
assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1286 |
assumes even: "\<And>x. f (- x) = f x" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1287 |
shows "has_bochner_integral lborel f (2 *\<^sub>R x)" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1288 |
proof - |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1289 |
have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1290 |
by (auto split: split_indicator) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1291 |
have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) x" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1292 |
by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1293 |
(auto simp: indicator even f) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1294 |
with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + x)" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1295 |
by (rule has_bochner_integral_add) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1296 |
then have "has_bochner_integral lborel f (x + x)" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1297 |
by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1298 |
(auto split: split_indicator) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1299 |
then show ?thesis |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1300 |
by (simp add: scaleR_2) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1301 |
qed |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1302 |
|
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1303 |
lemma has_bochner_integral_odd_function: |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1304 |
fixes f :: "real \<Rightarrow> 'a :: {banach, second_countable_topology}" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1305 |
assumes f: "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x) x" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1306 |
assumes odd: "\<And>x. f (- x) = - f x" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1307 |
shows "has_bochner_integral lborel f 0" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1308 |
proof - |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1309 |
have indicator: "\<And>x::real. indicator {..0} (- x) = indicator {0..} x" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1310 |
by (auto split: split_indicator) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1311 |
have "has_bochner_integral lborel (\<lambda>x. - indicator {.. 0} x *\<^sub>R f x) x" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1312 |
by (subst lborel_has_bochner_integral_real_affine_iff[where c="-1" and t=0]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1313 |
(auto simp: indicator odd f) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1314 |
from has_bochner_integral_minus[OF this] |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1315 |
have "has_bochner_integral lborel (\<lambda>x. indicator {.. 0} x *\<^sub>R f x) (- x)" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1316 |
by simp |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1317 |
with f have "has_bochner_integral lborel (\<lambda>x. indicator {0..} x *\<^sub>R f x + indicator {.. 0} x *\<^sub>R f x) (x + - x)" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1318 |
by (rule has_bochner_integral_add) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1319 |
then have "has_bochner_integral lborel f (x + - x)" |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1320 |
by (rule has_bochner_integral_discrete_difference[where X="{0}", THEN iffD1, rotated 4]) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1321 |
(auto split: split_indicator) |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1322 |
then show ?thesis |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1323 |
by simp |
0ddb5b755cdc
moved lemmas from the proof of the Central Limit Theorem by Jeremy Avigad and Luke Serafin
hoelzl
parents:
57235
diff
changeset
|
1324 |
qed |
57235
b0b9a10e4bf4
properties of Erlang and exponentially distributed random variables (by Sudeep Kanav)
hoelzl
parents:
57166
diff
changeset
|
1325 |
|
38656 | 1326 |
end |