| author | nipkow | 
| Wed, 17 Apr 2013 21:23:35 +0200 | |
| changeset 51712 | 30624dab6054 | 
| parent 51478 | 270b21f3ae0a | 
| child 53015 | a1119cf551e8 | 
| permissions | -rw-r--r-- | 
| 42067 | 1  | 
(* Title: HOL/Probability/Lebesgue_Measure.thy  | 
2  | 
Author: Johannes Hölzl, TU München  | 
|
3  | 
Author: Robert Himmelmann, TU München  | 
|
4  | 
*)  | 
|
5  | 
||
| 38656 | 6  | 
header {* Lebsegue measure *}
 | 
| 42067 | 7  | 
|
| 38656 | 8  | 
theory Lebesgue_Measure  | 
| 
42146
 
5b52c6a9c627
split Product_Measure into Binary_Product_Measure and Finite_Product_Measure
 
hoelzl 
parents: 
42067 
diff
changeset
 | 
9  | 
imports Finite_Product_Measure  | 
| 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"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50418 
diff
changeset
 | 
51  | 
unfolding cube_def subset_interval by (simp add: setsum_negf ex_in_conv)  | 
| 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"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50418 
diff
changeset
 | 
54  | 
apply (simp add: cube_def subset_eq mem_interval setsum_negf eucl_le[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
 | 
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)"  | 
| 
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
 | 
70  | 
unfolding cube_def subset_interval 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"  | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50418 
diff
changeset
 | 
84  | 
unfolding cube_def inter_interval by (rule has_integral_const)  | 
| 
 
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))"  | 
172  | 
proof (intro arg_cong[where f="SUPR 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 }  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
197  | 
ultimately show "(\<lambda>m. \<Sum>x = 0..<m. ?m n x) ----> ?M n UNIV"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
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"  | 
|
| 41654 | 225  | 
using assms by (force simp: 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)  | 
| 
 
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
 | 
234  | 
(auto simp: cube_def negligible_def)  | 
| 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  | 
|
384  | 
    have integrable: "(indicator A :: _\<Rightarrow>real) integrable_on {a..b}"
 | 
|
385  | 
by (intro integrable_on_subinterval has_integral_integrable) (auto intro: *)  | 
|
386  | 
    then have "integral {a..b} (indicator A) \<le> (integral UNIV (indicator A) :: real)"
 | 
|
| 47694 | 387  | 
using * by (auto intro!: integral_subset_le)  | 
| 41654 | 388  | 
    moreover have "(0::real) \<le> integral {a..b} (indicator A)"
 | 
389  | 
using integrable by (auto intro!: integral_nonneg)  | 
|
390  | 
    ultimately have "integral {a..b} (indicator A) = (0::real)"
 | 
|
391  | 
using integral_unique[OF *] by auto  | 
|
392  | 
    then show "(indicator A has_integral (0::real)) {a..b}"
 | 
|
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]  | 
| 
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
 | 
415  | 
by (auto simp: cube_def content_closed_interval_cases setprod_constant setsum_negf)  | 
| 
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"
 | 
|
| 46905 | 426  | 
unfolding integrable_indicator_UNIV by (simp add: integrable_const indicator_def [abs_def])  | 
| 41654 | 427  | 
from lmeasure_eq_integral[OF this] show ?thesis unfolding integral_indicator_UNIV  | 
| 46905 | 428  | 
by (simp add: indicator_def [abs_def])  | 
| 40859 | 429  | 
qed  | 
430  | 
||
431  | 
lemma lmeasure_singleton[simp]:  | 
|
| 47694 | 432  | 
  fixes a :: "'a::ordered_euclidean_space" shows "emeasure lebesgue {a} = 0"
 | 
| 41654 | 433  | 
using lmeasure_atLeastAtMost[of a a] by simp  | 
| 40859 | 434  | 
|
| 49777 | 435  | 
lemma AE_lebesgue_singleton:  | 
436  | 
fixes a :: "'a::ordered_euclidean_space" shows "AE x in lebesgue. x \<noteq> a"  | 
|
437  | 
  by (rule AE_I[where N="{a}"]) auto
 | 
|
438  | 
||
| 40859 | 439  | 
declare content_real[simp]  | 
440  | 
||
441  | 
lemma  | 
|
442  | 
fixes a b :: real  | 
|
443  | 
shows lmeasure_real_greaterThanAtMost[simp]:  | 
|
| 47694 | 444  | 
    "emeasure lebesgue {a <.. b} = ereal (if a \<le> b then b - a else 0)"
 | 
| 49777 | 445  | 
proof -  | 
446  | 
  have "emeasure lebesgue {a <.. b} = emeasure lebesgue {a .. b}"
 | 
|
447  | 
using AE_lebesgue_singleton[of a]  | 
|
448  | 
by (intro emeasure_eq_AE) auto  | 
|
| 40859 | 449  | 
then show ?thesis by auto  | 
| 49777 | 450  | 
qed  | 
| 40859 | 451  | 
|
452  | 
lemma  | 
|
453  | 
fixes a b :: real  | 
|
454  | 
shows lmeasure_real_atLeastLessThan[simp]:  | 
|
| 47694 | 455  | 
    "emeasure lebesgue {a ..< b} = ereal (if a \<le> b then b - a else 0)"
 | 
| 49777 | 456  | 
proof -  | 
457  | 
  have "emeasure lebesgue {a ..< b} = emeasure lebesgue {a .. b}"
 | 
|
458  | 
using AE_lebesgue_singleton[of b]  | 
|
459  | 
by (intro emeasure_eq_AE) auto  | 
|
| 41654 | 460  | 
then show ?thesis by auto  | 
| 49777 | 461  | 
qed  | 
| 41654 | 462  | 
|
463  | 
lemma  | 
|
464  | 
fixes a b :: real  | 
|
465  | 
shows lmeasure_real_greaterThanLessThan[simp]:  | 
|
| 47694 | 466  | 
    "emeasure lebesgue {a <..< b} = ereal (if a \<le> b then b - a else 0)"
 | 
| 49777 | 467  | 
proof -  | 
468  | 
  have "emeasure lebesgue {a <..< b} = emeasure lebesgue {a .. b}"
 | 
|
469  | 
using AE_lebesgue_singleton[of a] AE_lebesgue_singleton[of b]  | 
|
470  | 
by (intro emeasure_eq_AE) auto  | 
|
| 40859 | 471  | 
then show ?thesis by auto  | 
| 49777 | 472  | 
qed  | 
| 40859 | 473  | 
|
| 
41706
 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 
hoelzl 
parents: 
41704 
diff
changeset
 | 
474  | 
subsection {* Lebesgue-Borel measure *}
 | 
| 
 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 
hoelzl 
parents: 
41704 
diff
changeset
 | 
475  | 
|
| 47694 | 476  | 
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
 | 
477  | 
|
| 
 
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
 | 
478  | 
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
 | 
479  | 
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
 | 
480  | 
and sets_lborel[simp]: "sets lborel = sets borel"  | 
| 47694 | 481  | 
and measurable_lborel1[simp]: "measurable lborel = measurable borel"  | 
482  | 
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
 | 
483  | 
using sets.sigma_sets_eq[of borel]  | 
| 47694 | 484  | 
by (auto simp add: lborel_def measurable_def[abs_def])  | 
| 40859 | 485  | 
|
| 47694 | 486  | 
lemma emeasure_lborel[simp]: "A \<in> sets borel \<Longrightarrow> emeasure lborel A = emeasure lebesgue A"  | 
487  | 
by (rule emeasure_measure_of[OF lborel_def])  | 
|
488  | 
(auto simp: positive_def emeasure_nonneg countably_additive_def intro!: suminf_emeasure)  | 
|
| 40859 | 489  | 
|
| 
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
 | 
490  | 
interpretation lborel: sigma_finite_measure lborel  | 
| 47694 | 491  | 
proof (default, intro conjI exI[of _ "\<lambda>n. cube n"])  | 
492  | 
show "range cube \<subseteq> sets lborel" by (auto intro: borel_closed)  | 
|
493  | 
  { fix x :: 'a have "\<exists>n. x\<in>cube n" using mem_big_cube by auto }
 | 
|
494  | 
then show "(\<Union>i. cube i) = (space lborel :: 'a set)" using mem_big_cube by auto  | 
|
495  | 
show "\<forall>i. emeasure lborel (cube i) \<noteq> \<infinity>" by (simp add: cube_def)  | 
|
496  | 
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
 | 
497  | 
|
| 
 
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
 | 
498  | 
interpretation lebesgue: sigma_finite_measure lebesgue  | 
| 40859 | 499  | 
proof  | 
| 47694 | 500  | 
from lborel.sigma_finite guess A :: "nat \<Rightarrow> 'a set" ..  | 
501  | 
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>)"  | 
|
502  | 
by (intro exI[of _ A]) (auto simp: subset_eq)  | 
|
| 40859 | 503  | 
qed  | 
504  | 
||
| 49777 | 505  | 
lemma Int_stable_atLeastAtMost:  | 
506  | 
fixes x::"'a::ordered_euclidean_space"  | 
|
507  | 
  shows "Int_stable (range (\<lambda>(a, b::'a). {a..b}))"
 | 
|
508  | 
by (auto simp: inter_interval Int_stable_def)  | 
|
509  | 
||
510  | 
lemma lborel_eqI:  | 
|
511  | 
fixes M :: "'a::ordered_euclidean_space measure"  | 
|
512  | 
  assumes emeasure_eq: "\<And>a b. emeasure M {a .. b} = content {a .. b}"
 | 
|
513  | 
assumes sets_eq: "sets M = sets borel"  | 
|
514  | 
shows "lborel = M"  | 
|
515  | 
proof (rule measure_eqI_generator_eq[OF Int_stable_atLeastAtMost])  | 
|
516  | 
  let ?P = "\<Pi>\<^isub>M i\<in>{..<DIM('a::ordered_euclidean_space)}. lborel"
 | 
|
517  | 
  let ?E = "range (\<lambda>(a, b). {a..b} :: 'a set)"
 | 
|
518  | 
show "?E \<subseteq> Pow UNIV" "sets lborel = sigma_sets UNIV ?E" "sets M = sigma_sets UNIV ?E"  | 
|
519  | 
by (simp_all add: borel_eq_atLeastAtMost sets_eq)  | 
|
520  | 
||
521  | 
show "range cube \<subseteq> ?E" unfolding cube_def [abs_def] by auto  | 
|
522  | 
  { fix x :: 'a have "\<exists>n. x \<in> cube n" using mem_big_cube[of x] by fastforce }
 | 
|
523  | 
then show "(\<Union>i. cube i :: 'a set) = UNIV" by auto  | 
|
524  | 
||
525  | 
  { fix i show "emeasure lborel (cube i) \<noteq> \<infinity>" unfolding cube_def by auto }
 | 
|
526  | 
  { fix X assume "X \<in> ?E" then show "emeasure lborel X = emeasure M X"
 | 
|
527  | 
by (auto simp: emeasure_eq) }  | 
|
528  | 
qed  | 
|
529  | 
||
530  | 
lemma lebesgue_real_affine:  | 
|
531  | 
fixes c :: real assumes "c \<noteq> 0"  | 
|
532  | 
shows "lborel = density (distr lborel borel (\<lambda>x. t + c * x)) (\<lambda>_. \<bar>c\<bar>)" (is "_ = ?D")  | 
|
533  | 
proof (rule lborel_eqI)  | 
|
534  | 
  fix a b show "emeasure ?D {a..b} = content {a .. b}"
 | 
|
535  | 
proof cases  | 
|
536  | 
assume "0 < c"  | 
|
537  | 
    then have "(\<lambda>x. t + c * x) -` {a..b} = {(a - t) / c .. (b - t) / c}"
 | 
|
538  | 
by (auto simp: field_simps)  | 
|
539  | 
with `0 < c` show ?thesis  | 
|
540  | 
by (cases "a \<le> b")  | 
|
541  | 
(auto simp: field_simps emeasure_density positive_integral_distr positive_integral_cmult  | 
|
542  | 
borel_measurable_indicator' emeasure_distr)  | 
|
543  | 
next  | 
|
544  | 
assume "\<not> 0 < c" with `c \<noteq> 0` have "c < 0" by auto  | 
|
545  | 
    then have *: "(\<lambda>x. t + c * x) -` {a..b} = {(b - t) / c .. (a - t) / c}"
 | 
|
546  | 
by (auto simp: field_simps)  | 
|
547  | 
with `c < 0` show ?thesis  | 
|
548  | 
by (cases "a \<le> b")  | 
|
549  | 
(auto simp: field_simps emeasure_density positive_integral_distr  | 
|
550  | 
positive_integral_cmult borel_measurable_indicator' emeasure_distr)  | 
|
551  | 
qed  | 
|
552  | 
qed simp  | 
|
553  | 
||
554  | 
lemma lebesgue_integral_real_affine:  | 
|
555  | 
fixes c :: real assumes c: "c \<noteq> 0" and f: "f \<in> borel_measurable borel"  | 
|
556  | 
shows "(\<integral> x. f x \<partial> lborel) = \<bar>c\<bar> * (\<integral> x. f (t + c * x) \<partial>lborel)"  | 
|
557  | 
by (subst lebesgue_real_affine[OF c, of t])  | 
|
558  | 
(simp add: f integral_density integral_distr lebesgue_integral_cmult)  | 
|
559  | 
||
| 
41706
 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 
hoelzl 
parents: 
41704 
diff
changeset
 | 
560  | 
subsection {* Lebesgue integrable implies Gauge integrable *}
 | 
| 
 
a207a858d1f6
prefer p2e before e2p; use measure_unique_Int_stable_vimage;
 
hoelzl 
parents: 
41704 
diff
changeset
 | 
561  | 
|
| 40859 | 562  | 
lemma simple_function_has_integral:  | 
| 43920 | 563  | 
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
 | 
564  | 
assumes f:"simple_function lebesgue f"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
565  | 
  and f':"range f \<subseteq> {0..<\<infinity>}"
 | 
| 47694 | 566  | 
  and om:"\<And>x. x \<in> range f \<Longrightarrow> emeasure lebesgue (f -` {x} \<inter> UNIV) = \<infinity> \<Longrightarrow> x = 0"
 | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41661 
diff
changeset
 | 
567  | 
shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
568  | 
unfolding simple_integral_def space_lebesgue  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
569  | 
proof (subst lebesgue_simple_function_indicator)  | 
| 47694 | 570  | 
  let ?M = "\<lambda>x. emeasure lebesgue (f -` {x} \<inter> UNIV)"
 | 
| 46731 | 571  | 
  let ?F = "\<lambda>x. indicator (f -` {x})"
 | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
572  | 
  { fix x y assume "y \<in> range f"
 | 
| 43920 | 573  | 
from subsetD[OF f' this] have "y * ?F y x = ereal (real y * ?F y x)"  | 
574  | 
by (cases rule: ereal2_cases[of y "?F y x"])  | 
|
575  | 
(auto simp: indicator_def one_ereal_def split: split_if_asm) }  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
576  | 
moreover  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
577  | 
  { fix x assume x: "x\<in>range f"
 | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
578  | 
have "x * ?M x = real x * real (?M x)"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
579  | 
proof cases  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
580  | 
assume "x \<noteq> 0" with om[OF x] have "?M x \<noteq> \<infinity>" by auto  | 
| 47694 | 581  | 
with subsetD[OF f' x] f[THEN simple_functionD(2)] show ?thesis  | 
| 43920 | 582  | 
by (cases rule: ereal2_cases[of x "?M x"]) auto  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
583  | 
qed simp }  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
584  | 
ultimately  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
585  | 
have "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV \<longleftrightarrow>  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
586  | 
((\<lambda>x. \<Sum>y\<in>range f. real y * ?F y x) has_integral (\<Sum>x\<in>range f. real x * real (?M x))) UNIV"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
587  | 
by simp  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
588  | 
also have \<dots>  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
589  | 
proof (intro has_integral_setsum has_integral_cmult_real lmeasure_finite_has_integral  | 
| 47694 | 590  | 
real_of_ereal_pos emeasure_nonneg ballI)  | 
591  | 
    show *: "finite (range f)" "\<And>y. f -` {y} \<in> sets lebesgue"
 | 
|
592  | 
using simple_functionD[OF f] by auto  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
593  | 
fix y assume "real y \<noteq> 0" "y \<in> range f"  | 
| 47694 | 594  | 
    with * om[OF this(2)] show "emeasure lebesgue (f -` {y}) = ereal (real (?M y))"
 | 
| 43920 | 595  | 
by (auto simp: ereal_real)  | 
| 41654 | 596  | 
qed  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
597  | 
finally show "((\<lambda>x. real (\<Sum>y\<in>range f. y * ?F y x)) has_integral real (\<Sum>x\<in>range f. x * ?M x)) UNIV" .  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
598  | 
qed fact  | 
| 40859 | 599  | 
|
600  | 
lemma simple_function_has_integral':  | 
|
| 43920 | 601  | 
fixes f::"'a::ordered_euclidean_space \<Rightarrow> ereal"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
602  | 
assumes f: "simple_function lebesgue f" "\<And>x. 0 \<le> f x"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
603  | 
and i: "integral\<^isup>S lebesgue f \<noteq> \<infinity>"  | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41661 
diff
changeset
 | 
604  | 
shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>S lebesgue f))) UNIV"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
605  | 
proof -  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
606  | 
  let ?f = "\<lambda>x. if x \<in> f -` {\<infinity>} then 0 else f x"
 | 
| 47694 | 607  | 
note f(1)[THEN simple_functionD(2)]  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
608  | 
then have [simp, intro]: "\<And>X. f -` X \<in> sets lebesgue" by auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
609  | 
have f': "simple_function lebesgue ?f"  | 
| 47694 | 610  | 
using f by (intro simple_function_If_set) auto  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
611  | 
  have rng: "range ?f \<subseteq> {0..<\<infinity>}" using f by auto
 | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
612  | 
have "AE x in lebesgue. f x = ?f x"  | 
| 47694 | 613  | 
using simple_integral_PInf[OF f i]  | 
614  | 
    by (intro AE_I[where N="f -` {\<infinity>} \<inter> space lebesgue"]) auto
 | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
615  | 
from f(1) f' this have eq: "integral\<^isup>S lebesgue f = integral\<^isup>S lebesgue ?f"  | 
| 47694 | 616  | 
by (rule simple_integral_cong_AE)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
617  | 
have real_eq: "\<And>x. real (f x) = real (?f x)" by auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
618  | 
|
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
619  | 
show ?thesis  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
620  | 
unfolding eq real_eq  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
621  | 
proof (rule simple_function_has_integral[OF f' rng])  | 
| 47694 | 622  | 
    fix x assume x: "x \<in> range ?f" and inf: "emeasure lebesgue (?f -` {x} \<inter> UNIV) = \<infinity>"
 | 
623  | 
    have "x * emeasure lebesgue (?f -` {x} \<inter> UNIV) = (\<integral>\<^isup>S y. x * indicator (?f -` {x}) y \<partial>lebesgue)"
 | 
|
624  | 
using f'[THEN simple_functionD(2)]  | 
|
625  | 
by (simp add: simple_integral_cmult_indicator)  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
626  | 
also have "\<dots> \<le> integral\<^isup>S lebesgue f"  | 
| 47694 | 627  | 
using f'[THEN simple_functionD(2)] f  | 
628  | 
by (intro simple_integral_mono simple_function_mult simple_function_indicator)  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
629  | 
(auto split: split_indicator)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
630  | 
finally show "x = 0" unfolding inf using i subsetD[OF rng x] by (auto split: split_if_asm)  | 
| 40859 | 631  | 
qed  | 
632  | 
qed  | 
|
633  | 
||
634  | 
lemma positive_integral_has_integral:  | 
|
| 43920 | 635  | 
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
636  | 
  assumes f: "f \<in> borel_measurable lebesgue" "range f \<subseteq> {0..<\<infinity>}" "integral\<^isup>P lebesgue f \<noteq> \<infinity>"
 | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41661 
diff
changeset
 | 
637  | 
shows "((\<lambda>x. real (f x)) has_integral (real (integral\<^isup>P lebesgue f))) UNIV"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
638  | 
proof -  | 
| 47694 | 639  | 
from borel_measurable_implies_simple_function_sequence'[OF f(1)]  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
640  | 
guess u . note u = this  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
641  | 
have SUP_eq: "\<And>x. (SUP i. u i x) = f x"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
642  | 
using u(4) f(2)[THEN subsetD] by (auto split: split_max)  | 
| 46731 | 643  | 
let ?u = "\<lambda>i x. real (u i x)"  | 
| 47694 | 644  | 
note u_eq = positive_integral_eq_simple_integral[OF u(1,5), symmetric]  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
645  | 
  { fix i
 | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
646  | 
note u_eq  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
647  | 
also have "integral\<^isup>P lebesgue (u i) \<le> (\<integral>\<^isup>+x. max 0 (f x) \<partial>lebesgue)"  | 
| 47694 | 648  | 
by (intro positive_integral_mono) (auto intro: SUP_upper simp: u(4)[symmetric])  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
649  | 
finally have "integral\<^isup>S lebesgue (u i) \<noteq> \<infinity>"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
650  | 
unfolding positive_integral_max_0 using f by auto }  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
651  | 
note u_fin = this  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
652  | 
then have u_int: "\<And>i. (?u i has_integral real (integral\<^isup>S lebesgue (u i))) UNIV"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
653  | 
by (rule simple_function_has_integral'[OF u(1,5)])  | 
| 43920 | 654  | 
have "\<forall>x. \<exists>r\<ge>0. f x = ereal r"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
655  | 
proof  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
656  | 
fix x from f(2) have "0 \<le> f x" "f x \<noteq> \<infinity>" by (auto simp: subset_eq)  | 
| 43920 | 657  | 
then show "\<exists>r\<ge>0. f x = ereal r" by (cases "f x") auto  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
658  | 
qed  | 
| 43920 | 659  | 
from choice[OF this] obtain f' where f': "f = (\<lambda>x. ereal (f' x))" "\<And>x. 0 \<le> f' x" by auto  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
660  | 
|
| 43920 | 661  | 
have "\<forall>i. \<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
662  | 
proof  | 
| 43920 | 663  | 
fix i show "\<exists>r. \<forall>x. 0 \<le> r x \<and> u i x = ereal (r x)"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
664  | 
proof (intro choice allI)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
665  | 
fix i x have "u i x \<noteq> \<infinity>" using u(3)[of i] by (auto simp: image_iff) metis  | 
| 43920 | 666  | 
then show "\<exists>r\<ge>0. u i x = ereal r" using u(5)[of i x] by (cases "u i x") auto  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
667  | 
qed  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
668  | 
qed  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
669  | 
from choice[OF this] obtain u' where  | 
| 43920 | 670  | 
u': "u = (\<lambda>i x. ereal (u' i x))" "\<And>i x. 0 \<le> u' i x" by (auto simp: fun_eq_iff)  | 
| 40859 | 671  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
672  | 
have convergent: "f' integrable_on UNIV \<and>  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
673  | 
(\<lambda>k. integral UNIV (u' k)) ----> integral UNIV f'"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
674  | 
proof (intro monotone_convergence_increasing allI ballI)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
675  | 
show int: "\<And>k. (u' k) integrable_on UNIV"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
676  | 
using u_int unfolding integrable_on_def u' by auto  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
677  | 
show "\<And>k x. u' k x \<le> u' (Suc k) x" using u(2,3,5)  | 
| 43920 | 678  | 
by (auto simp: incseq_Suc_iff le_fun_def image_iff u' intro!: real_of_ereal_positive_mono)  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
679  | 
show "\<And>x. (\<lambda>k. u' k x) ----> f' x"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
680  | 
using SUP_eq u(2)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
681  | 
by (intro SUP_eq_LIMSEQ[THEN iffD1]) (auto simp: u' f' incseq_mono incseq_Suc_iff le_fun_def)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
682  | 
    show "bounded {integral UNIV (u' k)|k. True}"
 | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
683  | 
proof (safe intro!: bounded_realI)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
684  | 
fix k  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
685  | 
have "\<bar>integral UNIV (u' k)\<bar> = integral UNIV (u' k)"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
686  | 
by (intro abs_of_nonneg integral_nonneg int ballI u')  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
687  | 
also have "\<dots> = real (integral\<^isup>S lebesgue (u k))"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
688  | 
using u_int[THEN integral_unique] by (simp add: u')  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
689  | 
also have "\<dots> = real (integral\<^isup>P lebesgue (u k))"  | 
| 47694 | 690  | 
using positive_integral_eq_simple_integral[OF u(1,5)] by simp  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
691  | 
also have "\<dots> \<le> real (integral\<^isup>P lebesgue f)" using f  | 
| 47694 | 692  | 
by (auto intro!: real_of_ereal_positive_mono positive_integral_positive  | 
693  | 
positive_integral_mono SUP_upper simp: SUP_eq[symmetric])  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
694  | 
finally show "\<bar>integral UNIV (u' k)\<bar> \<le> real (integral\<^isup>P lebesgue f)" .  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
695  | 
qed  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
696  | 
qed  | 
| 40859 | 697  | 
|
| 43920 | 698  | 
have "integral\<^isup>P lebesgue f = ereal (integral UNIV f')"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
699  | 
proof (rule tendsto_unique[OF trivial_limit_sequentially])  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
700  | 
have "(\<lambda>i. integral\<^isup>S lebesgue (u i)) ----> (SUP i. integral\<^isup>P lebesgue (u i))"  | 
| 51000 | 701  | 
unfolding u_eq by (intro LIMSEQ_SUP incseq_positive_integral u)  | 
| 47694 | 702  | 
also note positive_integral_monotone_convergence_SUP  | 
703  | 
[OF u(2) borel_measurable_simple_function[OF u(1)] u(5), symmetric]  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
704  | 
finally show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> integral\<^isup>P lebesgue f"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
705  | 
unfolding SUP_eq .  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
706  | 
|
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
707  | 
    { fix k
 | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
708  | 
have "0 \<le> integral\<^isup>S lebesgue (u k)"  | 
| 47694 | 709  | 
using u by (auto intro!: simple_integral_positive)  | 
| 43920 | 710  | 
then have "integral\<^isup>S lebesgue (u k) = ereal (real (integral\<^isup>S lebesgue (u k)))"  | 
711  | 
using u_fin by (auto simp: ereal_real) }  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
712  | 
note * = this  | 
| 43920 | 713  | 
show "(\<lambda>k. integral\<^isup>S lebesgue (u k)) ----> ereal (integral UNIV f')"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
714  | 
using convergent using u_int[THEN integral_unique, symmetric]  | 
| 47694 | 715  | 
by (subst *) (simp add: u')  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
716  | 
qed  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
717  | 
then show ?thesis using convergent by (simp add: f' integrable_integral)  | 
| 40859 | 718  | 
qed  | 
719  | 
||
720  | 
lemma lebesgue_integral_has_integral:  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
721  | 
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
722  | 
assumes f: "integrable lebesgue f"  | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41661 
diff
changeset
 | 
723  | 
shows "(f has_integral (integral\<^isup>L lebesgue f)) UNIV"  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
724  | 
proof -  | 
| 43920 | 725  | 
let ?n = "\<lambda>x. real (ereal (max 0 (- f x)))" and ?p = "\<lambda>x. real (ereal (max 0 (f x)))"  | 
726  | 
have *: "f = (\<lambda>x. ?p x - ?n x)" by (auto simp del: ereal_max)  | 
|
| 47694 | 727  | 
  { fix f :: "'a \<Rightarrow> real" have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (\<integral>\<^isup>+ x. ereal (max 0 (f x)) \<partial>lebesgue)"
 | 
728  | 
by (intro positive_integral_cong_pos) (auto split: split_max) }  | 
|
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
729  | 
note eq = this  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
730  | 
show ?thesis  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
731  | 
unfolding lebesgue_integral_def  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
732  | 
apply (subst *)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
733  | 
apply (rule has_integral_sub)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
734  | 
unfolding eq[of f] eq[of "\<lambda>x. - f x"]  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
735  | 
apply (safe intro!: positive_integral_has_integral)  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
736  | 
using integrableD[OF f]  | 
| 43920 | 737  | 
by (auto simp: zero_ereal_def[symmetric] positive_integral_max_0 split: split_max  | 
| 47694 | 738  | 
intro!: measurable_If)  | 
| 40859 | 739  | 
qed  | 
740  | 
||
| 
47757
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
741  | 
lemma lebesgue_simple_integral_eq_borel:  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
742  | 
assumes f: "f \<in> borel_measurable borel"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
743  | 
shows "integral\<^isup>S lebesgue f = integral\<^isup>S lborel f"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
744  | 
using f[THEN measurable_sets]  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
745  | 
by (auto intro!: setsum_cong arg_cong2[where f="op *"] emeasure_lborel[symmetric]  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
746  | 
simp: simple_integral_def)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
747  | 
|
| 
41546
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
748  | 
lemma lebesgue_positive_integral_eq_borel:  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
749  | 
assumes f: "f \<in> borel_measurable borel"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
750  | 
shows "integral\<^isup>P lebesgue f = integral\<^isup>P lborel f"  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
751  | 
proof -  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
752  | 
from f have "integral\<^isup>P lebesgue (\<lambda>x. max 0 (f x)) = integral\<^isup>P lborel (\<lambda>x. max 0 (f x))"  | 
| 47694 | 753  | 
by (auto intro!: positive_integral_subalgebra[symmetric])  | 
| 
41981
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
754  | 
then show ?thesis unfolding positive_integral_max_0 .  | 
| 
 
cdf7693bbe08
reworked Probability theory: measures are not type restricted to positive extended reals
 
hoelzl 
parents: 
41831 
diff
changeset
 | 
755  | 
qed  | 
| 
41546
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
756  | 
|
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
757  | 
lemma lebesgue_integral_eq_borel:  | 
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
758  | 
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
 | 
759  | 
shows "integrable lebesgue f \<longleftrightarrow> integrable lborel f" (is ?P)  | 
| 
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41661 
diff
changeset
 | 
760  | 
and "integral\<^isup>L lebesgue f = integral\<^isup>L lborel f" (is ?I)  | 
| 
41546
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
761  | 
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
 | 
762  | 
have "sets lborel \<subseteq> sets lebesgue" by auto  | 
| 47694 | 763  | 
from integral_subalgebra[of f lborel, OF _ this _ _] assms  | 
| 
41546
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
764  | 
show ?P ?I by auto  | 
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
765  | 
qed  | 
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
766  | 
|
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
767  | 
lemma borel_integral_has_integral:  | 
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
768  | 
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
 | 
769  | 
assumes f:"integrable lborel f"  | 
| 
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41661 
diff
changeset
 | 
770  | 
shows "(f has_integral (integral\<^isup>L lborel f)) UNIV"  | 
| 
41546
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
771  | 
proof -  | 
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
772  | 
have borel: "f \<in> borel_measurable borel"  | 
| 
41689
 
3e39b0e730d6
the measure valuation is again part of the measure_space type, instead of an explicit parameter to the locale;
 
hoelzl 
parents: 
41661 
diff
changeset
 | 
773  | 
using f unfolding integrable_def by auto  | 
| 
41546
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
774  | 
from f show ?thesis  | 
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
775  | 
using lebesgue_integral_has_integral[of f]  | 
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
776  | 
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
 | 
777  | 
qed  | 
| 
 
2a12c23b7a34
integral on lebesgue measure is extension of integral on borel measure
 
hoelzl 
parents: 
41097 
diff
changeset
 | 
778  | 
|
| 49777 | 779  | 
lemma positive_integral_lebesgue_has_integral:  | 
| 
47757
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
780  | 
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"  | 
| 49777 | 781  | 
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
 | 
782  | 
assumes I: "(f has_integral I) UNIV"  | 
| 49777 | 783  | 
shows "(\<integral>\<^isup>+x. f x \<partial>lebesgue) = I"  | 
| 
47757
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
784  | 
proof -  | 
| 49777 | 785  | 
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
 | 
786  | 
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
 | 
787  | 
|
| 49777 | 788  | 
have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) = (SUP i. integral\<^isup>S lebesgue (F i))"  | 
| 
47757
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
789  | 
using F  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
790  | 
by (subst positive_integral_monotone_convergence_simple)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
791  | 
(simp_all add: positive_integral_max_0 simple_function_def)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
792  | 
also have "\<dots> \<le> ereal I"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
793  | 
proof (rule SUP_least)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
794  | 
fix i :: nat  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
795  | 
|
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
796  | 
    { fix z
 | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
797  | 
from F(4)[of z] have "F i z \<le> ereal (f z)"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
798  | 
by (metis SUP_upper UNIV_I ereal_max_0 min_max.sup_absorb2 nonneg)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
799  | 
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
 | 
800  | 
by (cases "F i z") simp_all }  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
801  | 
note F_bound = this  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
802  | 
|
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
803  | 
    { 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
 | 
804  | 
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
 | 
805  | 
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
 | 
806  | 
      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
 | 
807  | 
      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
 | 
808  | 
proof (rule dominated_convergence(1))  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
809  | 
fix n :: nat  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
810  | 
        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
 | 
811  | 
using x F(1)[of i]  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
812  | 
by (intro lebesgueD) (auto simp: simple_function_def)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
813  | 
then have cube: "?s n integrable_on cube n"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
814  | 
by (simp add: integrable_on_cmult_iff)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
815  | 
show "?s n integrable_on UNIV"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
816  | 
by (rule integrable_on_superset[OF _ _ cube]) auto  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
817  | 
next  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
818  | 
show "f integrable_on UNIV"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
819  | 
unfolding integrable_on_def using I by auto  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
820  | 
next  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
821  | 
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
 | 
822  | 
using nonneg F(5) by (auto split: split_indicator)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
823  | 
next  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
824  | 
        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
 | 
825  | 
proof  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
826  | 
fix z :: 'a  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
827  | 
from mem_big_cube[of z] guess j .  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
828  | 
          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
 | 
829  | 
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
 | 
830  | 
          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
 | 
831  | 
by (rule Lim_eventually)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
832  | 
qed  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
833  | 
qed  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
834  | 
      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
 | 
835  | 
by (simp add: integrable_on_cmult_iff) }  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
836  | 
note F_finite = lmeasure_finite[OF this]  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
837  | 
|
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
838  | 
have "((\<lambda>x. real (F i x)) has_integral real (integral\<^isup>S lebesgue (F i))) UNIV"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
839  | 
proof (rule simple_function_has_integral[of "F i"])  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
840  | 
show "simple_function lebesgue (F i)"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
841  | 
using F(1) by (simp add: simple_function_def)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
842  | 
      show "range (F i) \<subseteq> {0..<\<infinity>}"
 | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
843  | 
using F(3,5)[of i] by (auto simp: image_iff) metis  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
844  | 
next  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
845  | 
      fix x assume "x \<in> range (F i)" "emeasure lebesgue (F i -` {x} \<inter> UNIV) = \<infinity>"
 | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
846  | 
with F_finite[of x] show "x = 0" by auto  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
847  | 
qed  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
848  | 
from this I have "real (integral\<^isup>S lebesgue (F i)) \<le> I"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
849  | 
by (rule has_integral_le) (intro ballI F_bound)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
850  | 
moreover  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
851  | 
    { fix x assume x: "x \<in> range (F i)"
 | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
852  | 
with F(3,5)[of i] have "x = 0 \<or> (0 < x \<and> x \<noteq> \<infinity>)"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
853  | 
by (auto simp: image_iff le_less) metis  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
854  | 
      with F_finite[OF _ x] x have "x * emeasure lebesgue (F i -` {x} \<inter> UNIV) \<noteq> \<infinity>"
 | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
855  | 
by auto }  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
856  | 
then have "integral\<^isup>S lebesgue (F i) \<noteq> \<infinity>"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
857  | 
unfolding simple_integral_def setsum_Pinfty space_lebesgue by blast  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
858  | 
moreover have "0 \<le> integral\<^isup>S lebesgue (F i)"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
859  | 
using F(1,5) by (intro simple_integral_positive) (auto simp: simple_function_def)  | 
| 49777 | 860  | 
ultimately show "integral\<^isup>S lebesgue (F i) \<le> ereal I"  | 
861  | 
by (cases "integral\<^isup>S lebesgue (F i)") auto  | 
|
| 
47757
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
862  | 
qed  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
863  | 
also have "\<dots> < \<infinity>" by simp  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
864  | 
finally have finite: "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue) \<noteq> \<infinity>" by simp  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
865  | 
have borel: "(\<lambda>x. ereal (f x)) \<in> borel_measurable lebesgue"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
866  | 
using f_borel by (auto intro: borel_measurable_lebesgueI)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
867  | 
from positive_integral_has_integral[OF borel _ finite]  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
868  | 
have "(f has_integral real (\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue)) UNIV"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
869  | 
using nonneg by (simp add: subset_eq)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
870  | 
with I have "I = real (\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue)"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
871  | 
by (rule has_integral_unique)  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
872  | 
with finite positive_integral_positive[of _ "\<lambda>x. ereal (f x)"] show ?thesis  | 
| 49777 | 873  | 
by (cases "\<integral>\<^isup>+ x. ereal (f x) \<partial>lebesgue") auto  | 
| 
47757
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
874  | 
qed  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
875  | 
|
| 49777 | 876  | 
lemma has_integral_iff_positive_integral_lebesgue:  | 
877  | 
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"  | 
|
878  | 
assumes f: "f \<in> borel_measurable lebesgue" "\<And>x. 0 \<le> f x"  | 
|
879  | 
shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^isup>P lebesgue f = I"  | 
|
880  | 
using f positive_integral_lebesgue_has_integral[of f I] positive_integral_has_integral[of f]  | 
|
881  | 
by (auto simp: subset_eq)  | 
|
882  | 
||
883  | 
lemma has_integral_iff_positive_integral_lborel:  | 
|
| 
47757
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
884  | 
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
885  | 
assumes f: "f \<in> borel_measurable borel" "\<And>x. 0 \<le> f x"  | 
| 
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
886  | 
shows "(f has_integral I) UNIV \<longleftrightarrow> integral\<^isup>P lborel f = I"  | 
| 49777 | 887  | 
using assms  | 
888  | 
by (subst has_integral_iff_positive_integral_lebesgue)  | 
|
889  | 
(auto simp: borel_measurable_lebesgueI lebesgue_positive_integral_eq_borel)  | 
|
890  | 
||
891  | 
subsection {* Equivalence between product spaces and euclidean spaces *}
 | 
|
892  | 
||
| 
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
 | 
893  | 
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
 | 
894  | 
"e2p x = (\<lambda>i\<in>Basis. x \<bullet> i)"  | 
| 49777 | 895  | 
|
| 
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
 | 
896  | 
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
 | 
897  | 
"p2e x = (\<Sum>i\<in>Basis. x i *\<^sub>R i)"  | 
| 49777 | 898  | 
|
899  | 
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
 | 
900  | 
"x \<in> extensional Basis \<Longrightarrow> e2p (p2e x::'a::ordered_euclidean_space) = x"  | 
| 49777 | 901  | 
by (auto simp: fun_eq_iff extensional_def p2e_def e2p_def)  | 
902  | 
||
903  | 
lemma p2e_e2p[simp]:  | 
|
904  | 
"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
 | 
905  | 
by (auto simp: euclidean_eq_iff[where 'a='a] p2e_def e2p_def)  | 
| 49777 | 906  | 
|
907  | 
interpretation lborel_product: product_sigma_finite "\<lambda>x. lborel::real measure"  | 
|
908  | 
by default  | 
|
909  | 
||
| 
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
 | 
910  | 
interpretation lborel_space: finite_product_sigma_finite "\<lambda>x. lborel::real measure" "Basis"  | 
| 49777 | 911  | 
by default auto  | 
912  | 
||
913  | 
lemma sets_product_borel:  | 
|
914  | 
assumes I: "finite I"  | 
|
915  | 
  shows "sets (\<Pi>\<^isub>M i\<in>I. lborel) = sigma_sets (\<Pi>\<^isub>E i\<in>I. UNIV) { \<Pi>\<^isub>E i\<in>I. {..< x i :: real} | x. True}" (is "_ = ?G")
 | 
|
916  | 
proof (subst sigma_prod_algebra_sigma_eq[where S="\<lambda>_ i::nat. {..<real i}" and E="\<lambda>_. range lessThan", OF I])
 | 
|
917  | 
  show "sigma_sets (space (Pi\<^isub>M I (\<lambda>i. lborel))) {Pi\<^isub>E I F |F. \<forall>i\<in>I. F i \<in> range lessThan} = ?G"
 | 
|
918  | 
by (intro arg_cong2[where f=sigma_sets]) (auto simp: space_PiM image_iff bchoice_iff)  | 
|
| 
49779
 
1484b4b82855
remove incseq assumption from sigma_prod_algebra_sigma_eq
 
hoelzl 
parents: 
49777 
diff
changeset
 | 
919  | 
qed (auto simp: borel_eq_lessThan reals_Archimedean2)  | 
| 49777 | 920  | 
|
| 50003 | 921  | 
lemma measurable_e2p[measurable]:  | 
| 
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
 | 
922  | 
"e2p \<in> measurable (borel::'a::ordered_euclidean_space measure) (\<Pi>\<^isub>M (i::'a)\<in>Basis. (lborel :: real measure))"  | 
| 49777 | 923  | 
proof (rule measurable_sigma_sets[OF sets_product_borel])  | 
| 
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
 | 
924  | 
  fix A :: "('a \<Rightarrow> real) set" assume "A \<in> {\<Pi>\<^isub>E (i::'a)\<in>Basis. {..<x i} |x. True} "
 | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50418 
diff
changeset
 | 
925  | 
  then obtain x where  "A = (\<Pi>\<^isub>E (i::'a)\<in>Basis. {..<x i})" by auto
 | 
| 
 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 
hoelzl 
parents: 
50418 
diff
changeset
 | 
926  | 
  then have "e2p -` A = {..< (\<Sum>i\<in>Basis. x i *\<^sub>R i) :: 'a}"
 | 
| 
50123
 
69b35a75caf3
merge extensional dependent function space from FuncSet with the one in Finite_Product_Measure
 
hoelzl 
parents: 
50105 
diff
changeset
 | 
927  | 
using DIM_positive by (auto simp add: set_eq_iff e2p_def  | 
| 
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
 | 
928  | 
euclidean_eq_iff[where 'a='a] eucl_less[where 'a='a])  | 
| 49777 | 929  | 
then show "e2p -` A \<inter> space (borel::'a measure) \<in> sets borel" by simp  | 
930  | 
qed (auto simp: e2p_def)  | 
|
931  | 
||
| 50003 | 932  | 
(* FIXME: conversion in measurable prover *)  | 
| 50385 | 933  | 
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
 | 
934  | 
lemma lborelD[measurable (raw)]: "A \<in> sets borel \<Longrightarrow> A \<in> sets lborel" by simp  | 
|
| 50003 | 935  | 
|
936  | 
lemma measurable_p2e[measurable]:  | 
|
| 
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
 | 
937  | 
"p2e \<in> measurable (\<Pi>\<^isub>M (i::'a)\<in>Basis. (lborel :: real measure))  | 
| 49777 | 938  | 
(borel :: 'a::ordered_euclidean_space measure)"  | 
939  | 
(is "p2e \<in> measurable ?P _")  | 
|
940  | 
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
 | 
941  | 
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
 | 
942  | 
  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
 | 
943  | 
assume "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
 | 
944  | 
  then have "?A = (\<Pi>\<^isub>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
 | 
945  | 
using DIM_positive by (auto simp: space_PiM p2e_def PiE_def split: split_if_asm)  | 
| 49777 | 946  | 
then show "?A \<in> sets ?P"  | 
947  | 
by auto  | 
|
948  | 
qed  | 
|
949  | 
||
950  | 
lemma lborel_eq_lborel_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  | 
"(lborel :: 'a measure) = distr (\<Pi>\<^isub>M (i::'a::ordered_euclidean_space)\<in>Basis. lborel) borel p2e"  | 
| 49777 | 952  | 
(is "?B = ?D")  | 
953  | 
proof (rule lborel_eqI)  | 
|
954  | 
show "sets ?D = sets borel" 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
 | 
955  | 
let ?P = "(\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel)"  | 
| 49777 | 956  | 
fix a b :: '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
 | 
957  | 
  have *: "p2e -` {a .. b} \<inter> space ?P = (\<Pi>\<^isub>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
 | 
958  | 
by (auto simp: eucl_le[where 'a='a] p2e_def space_PiM PiE_def Pi_iff)  | 
| 49777 | 959  | 
  have "emeasure ?P (p2e -` {a..b} \<inter> space ?P) = content {a..b}"
 | 
960  | 
proof cases  | 
|
961  | 
    assume "{a..b} \<noteq> {}"
 | 
|
962  | 
then have "a \<le> b"  | 
|
963  | 
by (simp add: interval_ne_empty 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
 | 
964  | 
    then have "emeasure lborel {a..b} = (\<Prod>x\<in>Basis. emeasure lborel {a \<bullet> x .. b \<bullet> x})"
 | 
| 49777 | 965  | 
by (auto simp: content_closed_interval eucl_le[where 'a='a]  | 
966  | 
intro!: setprod_ereal[symmetric])  | 
|
967  | 
    also have "\<dots> = emeasure ?P (p2e -` {a..b} \<inter> space ?P)"
 | 
|
968  | 
unfolding * by (subst lborel_space.measure_times) auto  | 
|
969  | 
finally show ?thesis by simp  | 
|
970  | 
qed simp  | 
|
971  | 
  then show "emeasure ?D {a .. b} = content {a .. b}"
 | 
|
972  | 
by (simp add: emeasure_distr measurable_p2e)  | 
|
973  | 
qed  | 
|
974  | 
||
975  | 
lemma borel_fubini_positiv_integral:  | 
|
976  | 
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> ereal"  | 
|
977  | 
assumes f: "f \<in> borel_measurable borel"  | 
|
| 
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
 | 
978  | 
shows "integral\<^isup>P lborel f = \<integral>\<^isup>+x. f (p2e x) \<partial>(\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel)"  | 
| 49777 | 979  | 
by (subst lborel_eq_lborel_space) (simp add: positive_integral_distr measurable_p2e f)  | 
980  | 
||
981  | 
lemma borel_fubini_integrable:  | 
|
982  | 
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"  | 
|
| 
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
 | 
983  | 
shows "integrable lborel f \<longleftrightarrow> integrable (\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel) (\<lambda>x. f (p2e x))"  | 
| 49777 | 984  | 
(is "_ \<longleftrightarrow> integrable ?B ?f")  | 
985  | 
proof  | 
|
986  | 
assume "integrable lborel f"  | 
|
987  | 
moreover then have f: "f \<in> borel_measurable borel"  | 
|
988  | 
by auto  | 
|
989  | 
moreover with measurable_p2e  | 
|
990  | 
have "f \<circ> p2e \<in> borel_measurable ?B"  | 
|
991  | 
by (rule measurable_comp)  | 
|
992  | 
ultimately show "integrable ?B ?f"  | 
|
993  | 
by (simp add: comp_def borel_fubini_positiv_integral integrable_def)  | 
|
994  | 
next  | 
|
995  | 
assume "integrable ?B ?f"  | 
|
996  | 
moreover  | 
|
997  | 
then have "?f \<circ> e2p \<in> borel_measurable (borel::'a measure)"  | 
|
998  | 
by (auto intro!: measurable_e2p)  | 
|
999  | 
then have "f \<in> borel_measurable borel"  | 
|
1000  | 
by (simp cong: measurable_cong)  | 
|
1001  | 
ultimately show "integrable lborel f"  | 
|
1002  | 
by (simp add: borel_fubini_positiv_integral integrable_def)  | 
|
1003  | 
qed  | 
|
1004  | 
||
1005  | 
lemma borel_fubini:  | 
|
1006  | 
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"  | 
|
1007  | 
assumes f: "f \<in> borel_measurable borel"  | 
|
| 
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
 | 
1008  | 
shows "integral\<^isup>L lborel f = \<integral>x. f (p2e x) \<partial>((\<Pi>\<^isub>M (i::'a)\<in>Basis. lborel))"  | 
| 49777 | 1009  | 
using f by (simp add: borel_fubini_positiv_integral lebesgue_integral_def)  | 
| 
47757
 
5e6fe71e2390
equate positive Lebesgue integral and MV-Analysis' Gauge integral
 
hoelzl 
parents: 
47694 
diff
changeset
 | 
1010  | 
|
| 
50418
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1011  | 
lemma integrable_on_borel_integrable:  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1012  | 
fixes f :: "'a::ordered_euclidean_space \<Rightarrow> real"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1013  | 
assumes f_borel: "f \<in> borel_measurable borel" and nonneg: "\<And>x. 0 \<le> f x"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1014  | 
assumes f: "f integrable_on UNIV"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1015  | 
shows "integrable lborel f"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1016  | 
proof -  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1017  | 
have "(\<integral>\<^isup>+ x. ereal (f x) \<partial>lborel) \<noteq> \<infinity>"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1018  | 
using has_integral_iff_positive_integral_lborel[OF f_borel nonneg] f  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1019  | 
by (auto simp: integrable_on_def)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1020  | 
moreover have "(\<integral>\<^isup>+ x. ereal (- f x) \<partial>lborel) = 0"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1021  | 
using f_borel nonneg by (subst positive_integral_0_iff_AE) auto  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1022  | 
ultimately show ?thesis  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1023  | 
using f_borel by (auto simp: integrable_def)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1024  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1025  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1026  | 
subsection {* Fundamental Theorem of Calculus for the Lebesgue integral *}
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1027  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1028  | 
lemma borel_integrable_atLeastAtMost:  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1029  | 
fixes a b :: real  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1030  | 
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
 | 
1031  | 
  shows "integrable lborel (\<lambda>x. f x * indicator {a .. b} x)" (is "integrable _ ?f")
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1032  | 
proof cases  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1033  | 
assume "a \<le> b"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1034  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1035  | 
from isCont_Lb_Ub[OF `a \<le> b`, of f] f  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1036  | 
obtain M L where  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1037  | 
bounds: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> f x \<le> M" "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> L \<le> f x"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1038  | 
by metis  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1039  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1040  | 
show ?thesis  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1041  | 
proof (rule integrable_bound)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1042  | 
    show "integrable lborel (\<lambda>x. max \<bar>M\<bar> \<bar>L\<bar> * indicator {a..b} x)"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1043  | 
by (rule integral_cmul_indicator) simp_all  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1044  | 
    show "AE x in lborel. \<bar>?f x\<bar> \<le> max \<bar>M\<bar> \<bar>L\<bar> * indicator {a..b} x"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1045  | 
proof (rule AE_I2)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1046  | 
      fix x show "\<bar>?f x\<bar> \<le> max \<bar>M\<bar> \<bar>L\<bar> * indicator {a..b} x"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1047  | 
using bounds[of x] by (auto split: split_indicator)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1048  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1049  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1050  | 
    let ?g = "\<lambda>x. if x = a then f a else if x = b then f b else if x \<in> {a <..< b} then f x else 0"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1051  | 
    from f have "continuous_on {a <..< b} f"
 | 
| 
51478
 
270b21f3ae0a
move continuous and continuous_on to the HOL image; isCont is an abbreviation for continuous (at x) (isCont is now restricted to a T2 space)
 
hoelzl 
parents: 
51000 
diff
changeset
 | 
1052  | 
by (subst continuous_on_eq_continuous_at) auto  | 
| 
50418
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1053  | 
then have "?g \<in> borel_measurable borel"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1054  | 
      using borel_measurable_continuous_on_open[of "{a <..< b }" f "\<lambda>x. x" borel 0]
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1055  | 
by (auto intro!: measurable_If[where P="\<lambda>x. x = a"] measurable_If[where P="\<lambda>x. x = b"])  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1056  | 
also have "?g = ?f"  | 
| 
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
 | 
1057  | 
using `a \<le> b` by (intro ext) (auto split: split_indicator)  | 
| 
50418
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1058  | 
finally show "?f \<in> borel_measurable lborel"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1059  | 
by simp  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1060  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1061  | 
qed simp  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1062  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1063  | 
lemma integral_FTC_atLeastAtMost:  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1064  | 
fixes a b :: real  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1065  | 
assumes "a \<le> b"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1066  | 
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
 | 
1067  | 
and 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
 | 
1068  | 
  shows "integral\<^isup>L lborel (\<lambda>x. f x * indicator {a .. b} x) = F b - F a"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1069  | 
proof -  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1070  | 
  let ?f = "\<lambda>x. f x * indicator {a .. b} x"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1071  | 
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
 | 
1072  | 
using borel_integrable_atLeastAtMost[OF f]  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1073  | 
by (rule borel_integral_has_integral)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1074  | 
moreover  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1075  | 
  have "(f has_integral F b - F a) {a .. b}"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1076  | 
by (intro fundamental_theorem_of_calculus has_vector_derivative_withinI_DERIV ballI assms) auto  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1077  | 
  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
 | 
1078  | 
by (subst has_integral_eq_eq[where g=f]) auto  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1079  | 
then have "(?f has_integral F b - F a) UNIV"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1080  | 
    by (intro has_integral_on_superset[where t=UNIV and s="{a..b}"]) auto
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1081  | 
ultimately show "integral\<^isup>L lborel ?f = F b - F a"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1082  | 
by (rule has_integral_unique)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1083  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1084  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1085  | 
text {*
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1086  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1087  | 
For the positive integral we replace continuity with Borel-measurability.  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1088  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1089  | 
*}  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1090  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1091  | 
lemma positive_integral_FTC_atLeastAtMost:  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1092  | 
assumes f_borel: "f \<in> borel_measurable borel"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1093  | 
  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"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1094  | 
  shows "(\<integral>\<^isup>+x. f x * indicator {a .. b} x \<partial>lborel) = F b - F a"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1095  | 
proof -  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1096  | 
  have i: "(f has_integral F b - F a) {a..b}"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1097  | 
by (intro fundamental_theorem_of_calculus ballI has_vector_derivative_withinI_DERIV assms)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1098  | 
  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
 | 
1099  | 
by (rule has_integral_eq[OF _ i]) auto  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1100  | 
  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
 | 
1101  | 
by (rule has_integral_on_superset[OF _ _ i]) auto  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1102  | 
  then have "(\<integral>\<^isup>+x. ereal (f x * indicator {a .. b} x) \<partial>lborel) = F b - F a"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1103  | 
using f f_borel  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1104  | 
by (subst has_integral_iff_positive_integral_lborel[symmetric]) (auto split: split_indicator)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1105  | 
  also have "(\<integral>\<^isup>+x. ereal (f x * indicator {a .. b} x) \<partial>lborel) = (\<integral>\<^isup>+x. ereal (f x) * indicator {a .. b} x \<partial>lborel)"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1106  | 
by (auto intro!: positive_integral_cong simp: indicator_def)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1107  | 
finally show ?thesis by simp  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1108  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1109  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1110  | 
lemma positive_integral_FTC_atLeast:  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1111  | 
fixes f :: "real \<Rightarrow> real"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1112  | 
assumes f_borel: "f \<in> borel_measurable borel"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1113  | 
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
 | 
1114  | 
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
 | 
1115  | 
assumes lim: "(F ---> T) at_top"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1116  | 
  shows "(\<integral>\<^isup>+x. ereal (f x) * indicator {a ..} x \<partial>lborel) = T - F a"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1117  | 
proof -  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1118  | 
  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
 | 
1119  | 
  let ?fR = "\<lambda>x. ereal (f x) * indicator {a ..} x"
 | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1120  | 
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
 | 
1121  | 
proof (rule SUP_Lim_ereal)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1122  | 
show "\<And>x. incseq (\<lambda>i. ?f i x)"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1123  | 
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
 | 
1124  | 
|
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1125  | 
fix x  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1126  | 
from reals_Archimedean2[of "x - a"] guess n ..  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1127  | 
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
 | 
1128  | 
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
 | 
1129  | 
then show "(\<lambda>n. ?f n x) ----> ?fR x"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1130  | 
by (rule Lim_eventually)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1131  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1132  | 
then have "integral\<^isup>P lborel ?fR = (\<integral>\<^isup>+ x. (SUP i::nat. ?f i x) \<partial>lborel)"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1133  | 
by simp  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1134  | 
also have "\<dots> = (SUP i::nat. (\<integral>\<^isup>+ x. ?f i x \<partial>lborel))"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1135  | 
proof (rule positive_integral_monotone_convergence_SUP)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1136  | 
show "incseq ?f"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1137  | 
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
 | 
1138  | 
show "\<And>i. (?f i) \<in> borel_measurable lborel"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1139  | 
using f_borel by auto  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1140  | 
show "\<And>i x. 0 \<le> ?f i x"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1141  | 
using nonneg by (auto split: split_indicator)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1142  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1143  | 
also have "\<dots> = (SUP i::nat. F (a + real i) - F a)"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1144  | 
by (subst positive_integral_FTC_atLeastAtMost[OF f_borel f nonneg]) auto  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1145  | 
also have "\<dots> = T - F a"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1146  | 
proof (rule SUP_Lim_ereal)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1147  | 
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
 | 
1148  | 
proof (simp add: incseq_def, safe)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1149  | 
fix m n :: nat assume "m \<le> n"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1150  | 
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
 | 
1151  | 
by (intro DERIV_nonneg_imp_nondecreasing[where f=F])  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1152  | 
(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
 | 
1153  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1154  | 
have "(\<lambda>x. F (a + real x)) ----> T"  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1155  | 
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
 | 
1156  | 
apply (rule LIMSEQ_const_iff[THEN iffD2, OF refl])  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1157  | 
apply (rule filterlim_real_sequentially)  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1158  | 
done  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1159  | 
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
 | 
1160  | 
unfolding lim_ereal  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1161  | 
by (intro tendsto_diff) auto  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1162  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1163  | 
finally show ?thesis .  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1164  | 
qed  | 
| 
 
bd68cf816dd3
fundamental theorem of calculus for the Lebesgue integral
 
hoelzl 
parents: 
50385 
diff
changeset
 | 
1165  | 
|
| 38656 | 1166  | 
end  |