| author | wenzelm | 
| Wed, 20 Oct 2021 17:11:46 +0200 | |
| changeset 74560 | 5c8177fd1295 | 
| parent 74362 | 0135a0c77b64 | 
| child 77322 | 9c295f84d55f | 
| permissions | -rw-r--r-- | 
| 63627 | 1  | 
(* Title: HOL/Analysis/Set_Integral.thy  | 
| 63329 | 2  | 
Author: Jeremy Avigad (CMU), Johannes Hölzl (TUM), Luke Serafin (CMU)  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
3  | 
Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
4  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
5  | 
Notation and useful facts for working with integrals over a set.  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
6  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
7  | 
TODO: keep all these? Need unicode translations as well.  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
8  | 
*)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
9  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
10  | 
theory Set_Integral  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
11  | 
imports Radon_Nikodym  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
12  | 
begin  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
13  | 
|
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
14  | 
(*  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
15  | 
Notation  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
16  | 
*)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
17  | 
|
| 70136 | 18  | 
definition\<^marker>\<open>tag important\<close> "set_borel_measurable M A f \<equiv> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable M"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
19  | 
|
| 70136 | 20  | 
definition\<^marker>\<open>tag important\<close> "set_integrable M A f \<equiv> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
21  | 
|
| 70136 | 22  | 
definition\<^marker>\<open>tag important\<close> "set_lebesgue_integral M A f \<equiv> lebesgue_integral M (\<lambda>x. indicator A x *\<^sub>R f x)"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
23  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
24  | 
syntax  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
25  | 
"_ascii_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
26  | 
  ("(4LINT (_):(_)/|(_)./ _)" [0,60,110,61] 60)
 | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
27  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
28  | 
translations  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
29  | 
"LINT x:A|M. f" == "CONST set_lebesgue_integral M A (\<lambda>x. f)"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
30  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
31  | 
(*  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
32  | 
Notation for integration wrt lebesgue measure on the reals:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
33  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
34  | 
LBINT x. f  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
35  | 
LBINT x : A. f  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
36  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
37  | 
TODO: keep all these? Need unicode.  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
38  | 
*)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
39  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
40  | 
syntax  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
41  | 
"_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
42  | 
  ("(2LBINT _./ _)" [0,60] 60)
 | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
43  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
44  | 
syntax  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
45  | 
"_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> real"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
46  | 
  ("(3LBINT _:_./ _)" [0,60,61] 60)
 | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
47  | 
|
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
48  | 
(*  | 
| 
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
49  | 
Basic properties  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
50  | 
*)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
51  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
52  | 
(*  | 
| 61945 | 53  | 
lemma indicator_abs_eq: "\<And>A x. \<bar>indicator A x\<bar> = ((indicator A x) :: real)"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
54  | 
by (auto simp add: indicator_def)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
55  | 
*)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
56  | 
|
| 68721 | 57  | 
lemma set_integrable_cong:  | 
58  | 
assumes "M = M'" "A = A'" "\<And>x. x \<in> A \<Longrightarrow> f x = f' x"  | 
|
59  | 
shows "set_integrable M A f = set_integrable M' A' f'"  | 
|
60  | 
proof -  | 
|
61  | 
have "(\<lambda>x. indicator A x *\<^sub>R f x) = (\<lambda>x. indicator A' x *\<^sub>R f' x)"  | 
|
| 73536 | 62  | 
using assms by (auto simp: indicator_def of_bool_def)  | 
| 68721 | 63  | 
thus ?thesis by (simp add: set_integrable_def assms)  | 
64  | 
qed  | 
|
65  | 
||
| 62083 | 66  | 
lemma set_borel_measurable_sets:  | 
67  | 
fixes f :: "_ \<Rightarrow> _::real_normed_vector"  | 
|
68  | 
assumes "set_borel_measurable M X f" "B \<in> sets borel" "X \<in> sets M"  | 
|
69  | 
shows "f -` B \<inter> X \<in> sets M"  | 
|
70  | 
proof -  | 
|
71  | 
have "f \<in> borel_measurable (restrict_space M X)"  | 
|
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
72  | 
using assms unfolding set_borel_measurable_def by (subst borel_measurable_restrict_space_iff) auto  | 
| 62083 | 73  | 
then have "f -` B \<inter> space (restrict_space M X) \<in> sets (restrict_space M X)"  | 
74  | 
by (rule measurable_sets) fact  | 
|
75  | 
with \<open>X \<in> sets M\<close> show ?thesis  | 
|
76  | 
by (subst (asm) sets_restrict_space_iff) (auto simp: space_restrict_space)  | 
|
77  | 
qed  | 
|
78  | 
||
| 
67977
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67974 
diff
changeset
 | 
79  | 
lemma set_lebesgue_integral_zero [simp]: "set_lebesgue_integral M A (\<lambda>x. 0) = 0"  | 
| 
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67974 
diff
changeset
 | 
80  | 
by (auto simp: set_lebesgue_integral_def)  | 
| 
 
557ea2740125
Probability builds with new definitions
 
paulson <lp15@cam.ac.uk> 
parents: 
67974 
diff
changeset
 | 
81  | 
|
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
82  | 
lemma set_lebesgue_integral_cong:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
83  | 
assumes "A \<in> sets M" and "\<forall>x. x \<in> A \<longrightarrow> f x = g x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
84  | 
shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
85  | 
unfolding set_lebesgue_integral_def  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
86  | 
using assms  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
87  | 
by (metis indicator_simps(2) real_vector.scale_zero_left)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
88  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
89  | 
lemma set_lebesgue_integral_cong_AE:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
90  | 
assumes [measurable]: "A \<in> sets M" "f \<in> borel_measurable M" "g \<in> borel_measurable M"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
91  | 
assumes "AE x \<in> A in M. f x = g x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
92  | 
shows "LINT x:A|M. f x = LINT x:A|M. g x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
93  | 
proof-  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
94  | 
have "AE x in M. indicator A x *\<^sub>R f x = indicator A x *\<^sub>R g x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
95  | 
using assms by auto  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
96  | 
thus ?thesis  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
97  | 
unfolding set_lebesgue_integral_def by (intro integral_cong_AE) auto  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
98  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
99  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
100  | 
lemma set_integrable_cong_AE:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
101  | 
"f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>  | 
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
102  | 
AE x \<in> A in M. f x = g x \<Longrightarrow> A \<in> sets M \<Longrightarrow>  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
103  | 
set_integrable M A f = set_integrable M A g"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
104  | 
unfolding set_integrable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
105  | 
by (rule integrable_cong_AE) auto  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
106  | 
|
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
107  | 
lemma set_integrable_subset:  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
108  | 
  fixes M A B and f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
109  | 
assumes "set_integrable M A f" "B \<in> sets M" "B \<subseteq> A"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
110  | 
shows "set_integrable M B f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
111  | 
proof -  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
112  | 
have "set_integrable M B (\<lambda>x. indicator A x *\<^sub>R f x)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
113  | 
using assms integrable_mult_indicator set_integrable_def by blast  | 
| 61808 | 114  | 
with \<open>B \<subseteq> A\<close> show ?thesis  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
115  | 
unfolding set_integrable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
116  | 
by (simp add: indicator_inter_arith[symmetric] Int_absorb2)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
117  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
118  | 
|
| 67339 | 119  | 
lemma set_integrable_restrict_space:  | 
120  | 
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
|
121  | 
assumes f: "set_integrable M S f" and T: "T \<in> sets (restrict_space M S)"  | 
|
122  | 
shows "set_integrable M T f"  | 
|
123  | 
proof -  | 
|
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
124  | 
obtain T' where T_eq: "T = S \<inter> T'" and "T' \<in> sets M"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
125  | 
using T by (auto simp: sets_restrict_space)  | 
| 67339 | 126  | 
have \<open>integrable M (\<lambda>x. indicator T' x *\<^sub>R (indicator S x *\<^sub>R f x))\<close>  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
127  | 
using \<open>T' \<in> sets M\<close> f integrable_mult_indicator set_integrable_def by blast  | 
| 67339 | 128  | 
then show ?thesis  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
129  | 
unfolding set_integrable_def  | 
| 67339 | 130  | 
unfolding T_eq indicator_inter_arith by (simp add: ac_simps)  | 
131  | 
qed  | 
|
132  | 
||
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
133  | 
(* TODO: integral_cmul_indicator should be named set_integral_const *)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
134  | 
(* TODO: borel_integrable_atLeastAtMost should be named something like set_integrable_Icc_isCont *)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
135  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
136  | 
lemma set_integral_scaleR_right [simp]: "LINT t:A|M. a *\<^sub>R f t = a *\<^sub>R (LINT t:A|M. f t)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
137  | 
unfolding set_lebesgue_integral_def  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63627 
diff
changeset
 | 
138  | 
by (subst integral_scaleR_right[symmetric]) (auto intro!: Bochner_Integration.integral_cong)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
139  | 
|
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
140  | 
lemma set_integral_mult_right [simp]:  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
141  | 
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
142  | 
shows "LINT t:A|M. a * f t = a * (LINT t:A|M. f t)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
143  | 
unfolding set_lebesgue_integral_def  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63627 
diff
changeset
 | 
144  | 
by (subst integral_mult_right_zero[symmetric]) auto  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
145  | 
|
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
146  | 
lemma set_integral_mult_left [simp]:  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
147  | 
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
148  | 
shows "LINT t:A|M. f t * a = (LINT t:A|M. f t) * a"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
149  | 
unfolding set_lebesgue_integral_def  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63627 
diff
changeset
 | 
150  | 
by (subst integral_mult_left_zero[symmetric]) auto  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
151  | 
|
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
152  | 
lemma set_integral_divide_zero [simp]:  | 
| 
59867
 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 
haftmann 
parents: 
59358 
diff
changeset
 | 
153  | 
  fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
 | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
154  | 
shows "LINT t:A|M. f t / a = (LINT t:A|M. f t) / a"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
155  | 
unfolding set_lebesgue_integral_def  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63627 
diff
changeset
 | 
156  | 
by (subst integral_divide_zero[symmetric], intro Bochner_Integration.integral_cong)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
157  | 
(auto split: split_indicator)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
158  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
159  | 
lemma set_integrable_scaleR_right [simp, intro]:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
160  | 
shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. a *\<^sub>R f t)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
161  | 
unfolding set_integrable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
162  | 
unfolding scaleR_left_commute by (rule integrable_scaleR_right)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
163  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
164  | 
lemma set_integrable_scaleR_left [simp, intro]:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
165  | 
  fixes a :: "_ :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
166  | 
shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. f t *\<^sub>R a)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
167  | 
unfolding set_integrable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
168  | 
using integrable_scaleR_left[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
169  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
170  | 
lemma set_integrable_mult_right [simp, intro]:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
171  | 
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
172  | 
shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. a * f t)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
173  | 
unfolding set_integrable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
174  | 
using integrable_mult_right[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
175  | 
|
| 
70721
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
176  | 
lemma set_integrable_mult_right_iff [simp]:  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
177  | 
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
 | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
178  | 
assumes "a \<noteq> 0"  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
179  | 
shows "set_integrable M A (\<lambda>t. a * f t) \<longleftrightarrow> set_integrable M A f"  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
180  | 
proof  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
181  | 
assume "set_integrable M A (\<lambda>t. a * f t)"  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
182  | 
then have "set_integrable M A (\<lambda>t. 1/a * (a * f t))"  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
183  | 
using set_integrable_mult_right by blast  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
184  | 
then show "set_integrable M A f"  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
185  | 
using assms by auto  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
186  | 
qed auto  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
187  | 
|
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
188  | 
lemma set_integrable_mult_left [simp, intro]:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
189  | 
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
190  | 
shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. f t * a)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
191  | 
unfolding set_integrable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
192  | 
using integrable_mult_left[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
193  | 
|
| 
70721
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
194  | 
lemma set_integrable_mult_left_iff [simp]:  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
195  | 
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
 | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
196  | 
assumes "a \<noteq> 0"  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
197  | 
shows "set_integrable M A (\<lambda>t. f t * a) \<longleftrightarrow> set_integrable M A f"  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
198  | 
using assms by (subst set_integrable_mult_right_iff [symmetric]) (auto simp: mult.commute)  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
199  | 
|
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
200  | 
lemma set_integrable_divide [simp, intro]:  | 
| 
59867
 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 
haftmann 
parents: 
59358 
diff
changeset
 | 
201  | 
  fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
 | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
202  | 
assumes "a \<noteq> 0 \<Longrightarrow> set_integrable M A f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
203  | 
shows "set_integrable M A (\<lambda>t. f t / a)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
204  | 
proof -  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
205  | 
have "integrable M (\<lambda>x. indicator A x *\<^sub>R f x / a)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
206  | 
using assms unfolding set_integrable_def by (rule integrable_divide_zero)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
207  | 
also have "(\<lambda>x. indicator A x *\<^sub>R f x / a) = (\<lambda>x. indicator A x *\<^sub>R (f x / a))"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
208  | 
by (auto split: split_indicator)  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
209  | 
finally show ?thesis  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
210  | 
unfolding set_integrable_def .  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
211  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
212  | 
|
| 
70721
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
213  | 
lemma set_integrable_mult_divide_iff [simp]:  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
214  | 
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
 | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
215  | 
assumes "a \<noteq> 0"  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
216  | 
shows "set_integrable M A (\<lambda>t. f t / a) \<longleftrightarrow> set_integrable M A f"  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
217  | 
by (simp add: divide_inverse assms)  | 
| 
 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 
paulson <lp15@cam.ac.uk> 
parents: 
70365 
diff
changeset
 | 
218  | 
|
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
219  | 
lemma set_integral_add [simp, intro]:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
220  | 
  fixes f g :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
221  | 
assumes "set_integrable M A f" "set_integrable M A g"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
222  | 
shows "set_integrable M A (\<lambda>x. f x + g x)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
223  | 
and "LINT x:A|M. f x + g x = (LINT x:A|M. f x) + (LINT x:A|M. g x)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
224  | 
using assms unfolding set_integrable_def set_lebesgue_integral_def by (simp_all add: scaleR_add_right)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
225  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
226  | 
lemma set_integral_diff [simp, intro]:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
227  | 
assumes "set_integrable M A f" "set_integrable M A g"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
228  | 
shows "set_integrable M A (\<lambda>x. f x - g x)" and "LINT x:A|M. f x - g x =  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
229  | 
(LINT x:A|M. f x) - (LINT x:A|M. g x)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
230  | 
using assms unfolding set_integrable_def set_lebesgue_integral_def by (simp_all add: scaleR_diff_right)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
231  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
232  | 
lemma set_integral_uminus: "set_integrable M A f \<Longrightarrow> LINT x:A|M. - f x = - (LINT x:A|M. f x)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
233  | 
unfolding set_integrable_def set_lebesgue_integral_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
234  | 
by (subst integral_minus[symmetric]) simp_all  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
235  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
236  | 
lemma set_integral_complex_of_real:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
237  | 
"LINT x:A|M. complex_of_real (f x) = of_real (LINT x:A|M. f x)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
238  | 
unfolding set_lebesgue_integral_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
239  | 
by (subst integral_complex_of_real[symmetric])  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63627 
diff
changeset
 | 
240  | 
(auto intro!: Bochner_Integration.integral_cong split: split_indicator)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
241  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
242  | 
lemma set_integral_mono:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
243  | 
fixes f g :: "_ \<Rightarrow> real"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
244  | 
assumes "set_integrable M A f" "set_integrable M A g"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
245  | 
"\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
246  | 
shows "(LINT x:A|M. f x) \<le> (LINT x:A|M. g x)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
247  | 
using assms unfolding set_integrable_def set_lebesgue_integral_def  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
248  | 
by (auto intro: integral_mono split: split_indicator)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
249  | 
|
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
250  | 
lemma set_integral_mono_AE:  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
251  | 
fixes f g :: "_ \<Rightarrow> real"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
252  | 
assumes "set_integrable M A f" "set_integrable M A g"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
253  | 
"AE x \<in> A in M. f x \<le> g x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
254  | 
shows "(LINT x:A|M. f x) \<le> (LINT x:A|M. g x)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
255  | 
using assms unfolding set_integrable_def set_lebesgue_integral_def  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
256  | 
by (auto intro: integral_mono_AE split: split_indicator)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
257  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
258  | 
lemma set_integrable_abs: "set_integrable M A f \<Longrightarrow> set_integrable M A (\<lambda>x. \<bar>f x\<bar> :: real)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
259  | 
using integrable_abs[of M "\<lambda>x. f x * indicator A x"]unfolding set_integrable_def by (simp add: abs_mult ac_simps)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
260  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
261  | 
lemma set_integrable_abs_iff:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
262  | 
fixes f :: "_ \<Rightarrow> real"  | 
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
263  | 
shows "set_borel_measurable M A f \<Longrightarrow> set_integrable M A (\<lambda>x. \<bar>f x\<bar>) = set_integrable M A f"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
264  | 
unfolding set_integrable_def set_borel_measurable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
265  | 
by (subst (2) integrable_abs_iff[symmetric]) (simp_all add: abs_mult ac_simps)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
266  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
267  | 
lemma set_integrable_abs_iff':  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
268  | 
fixes f :: "_ \<Rightarrow> real"  | 
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
269  | 
shows "f \<in> borel_measurable M \<Longrightarrow> A \<in> sets M \<Longrightarrow>  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
270  | 
set_integrable M A (\<lambda>x. \<bar>f x\<bar>) = set_integrable M A f"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
271  | 
by (simp add: set_borel_measurable_def set_integrable_abs_iff)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
272  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
273  | 
lemma set_integrable_discrete_difference:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
274  | 
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
275  | 
assumes "countable X"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
276  | 
assumes diff: "(A - B) \<union> (B - A) \<subseteq> X"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
277  | 
  assumes "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
278  | 
shows "set_integrable M A f \<longleftrightarrow> set_integrable M B f"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
279  | 
unfolding set_integrable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
280  | 
proof (rule integrable_discrete_difference[where X=X])  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
281  | 
show "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
282  | 
using diff by (auto split: split_indicator)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
283  | 
qed fact+  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
284  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
285  | 
lemma set_integral_discrete_difference:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
286  | 
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
287  | 
assumes "countable X"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
288  | 
assumes diff: "(A - B) \<union> (B - A) \<subseteq> X"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
289  | 
  assumes "\<And>x. x \<in> X \<Longrightarrow> emeasure M {x} = 0" "\<And>x. x \<in> X \<Longrightarrow> {x} \<in> sets M"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
290  | 
shows "set_lebesgue_integral M A f = set_lebesgue_integral M B f"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
291  | 
unfolding set_lebesgue_integral_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
292  | 
proof (rule integral_discrete_difference[where X=X])  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
293  | 
show "\<And>x. x \<in> space M \<Longrightarrow> x \<notin> X \<Longrightarrow> indicator A x *\<^sub>R f x = indicator B x *\<^sub>R f x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
294  | 
using diff by (auto split: split_indicator)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
295  | 
qed fact+  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
296  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
297  | 
lemma set_integrable_Un:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
298  | 
  fixes f g :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
299  | 
assumes f_A: "set_integrable M A f" and f_B: "set_integrable M B f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
300  | 
and [measurable]: "A \<in> sets M" "B \<in> sets M"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
301  | 
shows "set_integrable M (A \<union> B) f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
302  | 
proof -  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
303  | 
have "set_integrable M (A - B) f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
304  | 
using f_A by (rule set_integrable_subset) auto  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
305  | 
with f_B have "integrable M (\<lambda>x. indicator (A - B) x *\<^sub>R f x + indicator B x *\<^sub>R f x)"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
306  | 
unfolding set_integrable_def using integrable_add by blast  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
307  | 
then show ?thesis  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
308  | 
unfolding set_integrable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
309  | 
by (rule integrable_cong[THEN iffD1, rotated 2]) (auto split: split_indicator)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
310  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
311  | 
|
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
312  | 
lemma set_integrable_empty [simp]: "set_integrable M {} f"
 | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
313  | 
by (auto simp: set_integrable_def)  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
314  | 
|
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
315  | 
lemma set_integrable_UN:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
316  | 
  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
317  | 
assumes "finite I" "\<And>i. i\<in>I \<Longrightarrow> set_integrable M (A i) f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
318  | 
"\<And>i. i\<in>I \<Longrightarrow> A i \<in> sets M"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
319  | 
shows "set_integrable M (\<Union>i\<in>I. A i) f"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
320  | 
using assms  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
321  | 
by (induct I) (auto simp: set_integrable_Un sets.finite_UN)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
322  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
323  | 
lemma set_integral_Un:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
324  | 
  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
325  | 
  assumes "A \<inter> B = {}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
326  | 
and "set_integrable M A f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
327  | 
and "set_integrable M B f"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
328  | 
shows "LINT x:A\<union>B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
329  | 
using assms  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
330  | 
unfolding set_integrable_def set_lebesgue_integral_def  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
331  | 
by (auto simp add: indicator_union_arith indicator_inter_arith[symmetric] scaleR_add_left)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
332  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
333  | 
lemma set_integral_cong_set:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
334  | 
  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
335  | 
assumes "set_borel_measurable M A f" "set_borel_measurable M B f"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
336  | 
and ae: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
337  | 
shows "LINT x:B|M. f x = LINT x:A|M. f x"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
338  | 
unfolding set_lebesgue_integral_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
339  | 
proof (rule integral_cong_AE)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
340  | 
show "AE x in M. indicator B x *\<^sub>R f x = indicator A x *\<^sub>R f x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
341  | 
using ae by (auto simp: subset_eq split: split_indicator)  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
342  | 
qed (use assms in \<open>auto simp: set_borel_measurable_def\<close>)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
343  | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
344  | 
proposition set_borel_measurable_subset:  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
345  | 
  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
346  | 
assumes [measurable]: "set_borel_measurable M A f" "B \<in> sets M" and "B \<subseteq> A"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
347  | 
shows "set_borel_measurable M B f"  | 
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
348  | 
proof-  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
349  | 
have "set_borel_measurable M B (\<lambda>x. indicator A x *\<^sub>R f x)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
350  | 
using assms unfolding set_borel_measurable_def  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
351  | 
using borel_measurable_indicator borel_measurable_scaleR by blast  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
352  | 
moreover have "(\<lambda>x. indicator B x *\<^sub>R indicator A x *\<^sub>R f x) = (\<lambda>x. indicator B x *\<^sub>R f x)"  | 
| 61808 | 353  | 
using \<open>B \<subseteq> A\<close> by (auto simp: fun_eq_iff split: split_indicator)  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
354  | 
ultimately show ?thesis  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
355  | 
unfolding set_borel_measurable_def by simp  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
356  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
357  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
358  | 
lemma set_integral_Un_AE:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
359  | 
  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
360  | 
assumes ae: "AE x in M. \<not> (x \<in> A \<and> x \<in> B)" and [measurable]: "A \<in> sets M" "B \<in> sets M"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
361  | 
and "set_integrable M A f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
362  | 
and "set_integrable M B f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
363  | 
shows "LINT x:A\<union>B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
364  | 
proof -  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
365  | 
have f: "set_integrable M (A \<union> B) f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
366  | 
by (intro set_integrable_Un assms)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
367  | 
then have f': "set_borel_measurable M (A \<union> B) f"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
368  | 
using integrable_iff_bounded set_borel_measurable_def set_integrable_def by blast  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
369  | 
have "LINT x:A\<union>B|M. f x = LINT x:(A - A \<inter> B) \<union> (B - A \<inter> B)|M. f x"  | 
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
370  | 
proof (rule set_integral_cong_set)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
371  | 
show "AE x in M. (x \<in> A - A \<inter> B \<union> (B - A \<inter> B)) = (x \<in> A \<union> B)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
372  | 
using ae by auto  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
373  | 
show "set_borel_measurable M (A - A \<inter> B \<union> (B - A \<inter> B)) f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
374  | 
using f' by (rule set_borel_measurable_subset) auto  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
375  | 
qed fact  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
376  | 
also have "\<dots> = (LINT x:(A - A \<inter> B)|M. f x) + (LINT x:(B - A \<inter> B)|M. f x)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
377  | 
by (auto intro!: set_integral_Un set_integrable_subset[OF f])  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
378  | 
also have "\<dots> = (LINT x:A|M. f x) + (LINT x:B|M. f x)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
379  | 
using ae  | 
| 67399 | 380  | 
by (intro arg_cong2[where f="(+)"] set_integral_cong_set)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
381  | 
(auto intro!: set_borel_measurable_subset[OF f'])  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
382  | 
finally show ?thesis .  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
383  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
384  | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
385  | 
lemma set_integral_finite_Union:  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
386  | 
  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
387  | 
assumes "finite I" "disjoint_family_on A I"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
388  | 
and "\<And>i. i \<in> I \<Longrightarrow> set_integrable M (A i) f" "\<And>i. i \<in> I \<Longrightarrow> A i \<in> sets M"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
389  | 
shows "(LINT x:(\<Union>i\<in>I. A i)|M. f x) = (\<Sum>i\<in>I. LINT x:A i|M. f x)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
390  | 
using assms  | 
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
391  | 
proof induction  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
392  | 
case (insert x F)  | 
| 69313 | 393  | 
  then have "A x \<inter> \<Union>(A ` F) = {}"
 | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
394  | 
by (meson disjoint_family_on_insert)  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
395  | 
with insert show ?case  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
396  | 
by (simp add: set_integral_Un set_integrable_Un set_integrable_UN disjoint_family_on_insert)  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
397  | 
qed (simp add: set_lebesgue_integral_def)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
398  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
399  | 
(* TODO: find a better name? *)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
400  | 
lemma pos_integrable_to_top:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
401  | 
fixes l::real  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
402  | 
assumes "\<And>i. A i \<in> sets M" "mono A"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
403  | 
assumes nneg: "\<And>x i. x \<in> A i \<Longrightarrow> 0 \<le> f x"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
404  | 
and intgbl: "\<And>i::nat. set_integrable M (A i) f"  | 
| 61969 | 405  | 
and lim: "(\<lambda>i::nat. LINT x:A i|M. f x) \<longlonglongrightarrow> l"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
406  | 
shows "set_integrable M (\<Union>i. A i) f"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
407  | 
unfolding set_integrable_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
408  | 
apply (rule integrable_monotone_convergence[where f = "\<lambda>i::nat. \<lambda>x. indicator (A i) x *\<^sub>R f x" and x = l])  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
409  | 
apply (rule intgbl [unfolded set_integrable_def])  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
410  | 
prefer 3 apply (rule lim [unfolded set_lebesgue_integral_def])  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
411  | 
apply (rule AE_I2)  | 
| 61808 | 412  | 
using \<open>mono A\<close> apply (auto simp: mono_def nneg split: split_indicator) []  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
413  | 
proof (rule AE_I2)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
414  | 
  { fix x assume "x \<in> space M"
 | 
| 61969 | 415  | 
show "(\<lambda>i. indicator (A i) x *\<^sub>R f x) \<longlonglongrightarrow> indicator (\<Union>i. A i) x *\<^sub>R f x"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
416  | 
proof cases  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
417  | 
assume "\<exists>i. x \<in> A i"  | 
| 74362 | 418  | 
then obtain i where "x \<in> A i" ..  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
419  | 
then have *: "eventually (\<lambda>i. x \<in> A i) sequentially"  | 
| 61808 | 420  | 
using \<open>x \<in> A i\<close> \<open>mono A\<close> by (auto simp: eventually_sequentially mono_def)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
421  | 
show ?thesis  | 
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
422  | 
apply (intro tendsto_eventually)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
423  | 
using *  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
424  | 
apply eventually_elim  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
425  | 
apply (auto split: split_indicator)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
426  | 
done  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
427  | 
qed auto }  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
428  | 
then show "(\<lambda>x. indicator (\<Union>i. A i) x *\<^sub>R f x) \<in> borel_measurable M"  | 
| 
62624
 
59ceeb6f3079
generalized some Borel measurable statements to support ennreal
 
hoelzl 
parents: 
62083 
diff
changeset
 | 
429  | 
apply (rule borel_measurable_LIMSEQ_real)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
430  | 
apply assumption  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
431  | 
using intgbl set_integrable_def by blast  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
432  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
433  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
434  | 
(* Proof from Royden Real Analysis, p. 91. *)  | 
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
435  | 
lemma lebesgue_integral_countable_add:  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
436  | 
  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
437  | 
assumes meas[intro]: "\<And>i::nat. A i \<in> sets M"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
438  | 
    and disj: "\<And>i j. i \<noteq> j \<Longrightarrow> A i \<inter> A j = {}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
439  | 
and intgbl: "set_integrable M (\<Union>i. A i) f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
440  | 
shows "LINT x:(\<Union>i. A i)|M. f x = (\<Sum>i. (LINT x:(A i)|M. f x))"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
441  | 
unfolding set_lebesgue_integral_def  | 
| 70136 | 442  | 
proof (subst integral_suminf[symmetric])  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
443  | 
show int_A: "integrable M (\<lambda>x. indicat_real (A i) x *\<^sub>R f x)" for i  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
444  | 
using intgbl unfolding set_integrable_def [symmetric]  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
445  | 
by (rule set_integrable_subset) auto  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
446  | 
  { fix x assume "x \<in> space M"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
447  | 
have "(\<lambda>i. indicator (A i) x *\<^sub>R f x) sums (indicator (\<Union>i. A i) x *\<^sub>R f x)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
448  | 
by (intro sums_scaleR_left indicator_sums) fact }  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
449  | 
note sums = this  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
450  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
451  | 
have norm_f: "\<And>i. set_integrable M (A i) (\<lambda>x. norm (f x))"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
452  | 
using int_A[THEN integrable_norm] unfolding set_integrable_def by auto  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
453  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
454  | 
show "AE x in M. summable (\<lambda>i. norm (indicator (A i) x *\<^sub>R f x))"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
455  | 
using disj by (intro AE_I2) (auto intro!: summable_mult2 sums_summable[OF indicator_sums])  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
456  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
457  | 
show "summable (\<lambda>i. LINT x|M. norm (indicator (A i) x *\<^sub>R f x))"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
458  | 
proof (rule summableI_nonneg_bounded)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
459  | 
fix n  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
460  | 
show "0 \<le> LINT x|M. norm (indicator (A n) x *\<^sub>R f x)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
461  | 
using norm_f by (auto intro!: integral_nonneg_AE)  | 
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
462  | 
|
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
463  | 
have "(\<Sum>i<n. LINT x|M. norm (indicator (A i) x *\<^sub>R f x)) = (\<Sum>i<n. LINT x:A i|M. norm (f x))"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
464  | 
by (simp add: abs_mult set_lebesgue_integral_def)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
465  | 
also have "\<dots> = set_lebesgue_integral M (\<Union>i<n. A i) (\<lambda>x. norm (f x))"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
466  | 
using norm_f  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
467  | 
by (subst set_integral_finite_Union) (auto simp: disjoint_family_on_def disj)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
468  | 
also have "\<dots> \<le> set_lebesgue_integral M (\<Union>i. A i) (\<lambda>x. norm (f x))"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
469  | 
using intgbl[unfolded set_integrable_def, THEN integrable_norm] norm_f  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
470  | 
unfolding set_lebesgue_integral_def set_integrable_def  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
471  | 
      apply (intro integral_mono set_integrable_UN[of "{..<n}", unfolded set_integrable_def])
 | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
472  | 
apply (auto split: split_indicator)  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
473  | 
done  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
474  | 
finally show "(\<Sum>i<n. LINT x|M. norm (indicator (A i) x *\<^sub>R f x)) \<le>  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
475  | 
set_lebesgue_integral M (\<Union>i. A i) (\<lambda>x. norm (f x))"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
476  | 
by simp  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
477  | 
qed  | 
| 69313 | 478  | 
show "LINT x|M. indicator (\<Union>(A ` UNIV)) x *\<^sub>R f x = LINT x|M. (\<Sum>i. indicator (A i) x *\<^sub>R f x)"  | 
| 
63886
 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 
hoelzl 
parents: 
63627 
diff
changeset
 | 
479  | 
apply (rule Bochner_Integration.integral_cong[OF refl])  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
480  | 
apply (subst suminf_scaleR_left[OF sums_summable[OF indicator_sums, OF disj], symmetric])  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
481  | 
using sums_unique[OF indicator_sums[OF disj]]  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
482  | 
apply auto  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
483  | 
done  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
484  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
485  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
486  | 
lemma set_integral_cont_up:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
487  | 
  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
488  | 
assumes [measurable]: "\<And>i. A i \<in> sets M" and A: "incseq A"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
489  | 
and intgbl: "set_integrable M (\<Union>i. A i) f"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
490  | 
shows "(\<lambda>i. LINT x:(A i)|M. f x) \<longlonglongrightarrow> LINT x:(\<Union>i. A i)|M. f x"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
491  | 
unfolding set_lebesgue_integral_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
492  | 
proof (intro integral_dominated_convergence[where w="\<lambda>x. indicator (\<Union>i. A i) x *\<^sub>R norm (f x)"])  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
493  | 
have int_A: "\<And>i. set_integrable M (A i) f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
494  | 
using intgbl by (rule set_integrable_subset) auto  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
495  | 
show "\<And>i. (\<lambda>x. indicator (A i) x *\<^sub>R f x) \<in> borel_measurable M"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
496  | 
using int_A integrable_iff_bounded set_integrable_def by blast  | 
| 69313 | 497  | 
show "(\<lambda>x. indicator (\<Union>(A ` UNIV)) x *\<^sub>R f x) \<in> borel_measurable M"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
498  | 
using integrable_iff_bounded intgbl set_integrable_def by blast  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
499  | 
show "integrable M (\<lambda>x. indicator (\<Union>i. A i) x *\<^sub>R norm (f x))"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
500  | 
using int_A intgbl integrable_norm unfolding set_integrable_def  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
501  | 
by fastforce  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
502  | 
  { fix x i assume "x \<in> A i"
 | 
| 61969 | 503  | 
with A have "(\<lambda>xa. indicator (A xa) x::real) \<longlonglongrightarrow> 1 \<longleftrightarrow> (\<lambda>xa. 1::real) \<longlonglongrightarrow> 1"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
504  | 
by (intro filterlim_cong refl)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
505  | 
(fastforce simp: eventually_sequentially incseq_def subset_eq intro!: exI[of _ i]) }  | 
| 61969 | 506  | 
then show "AE x in M. (\<lambda>i. indicator (A i) x *\<^sub>R f x) \<longlonglongrightarrow> indicator (\<Union>i. A i) x *\<^sub>R f x"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
507  | 
by (intro AE_I2 tendsto_intros) (auto split: split_indicator)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
508  | 
qed (auto split: split_indicator)  | 
| 
60615
 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 
paulson <lp15@cam.ac.uk> 
parents: 
59867 
diff
changeset
 | 
509  | 
|
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
510  | 
(* Can the int0 hypothesis be dropped? *)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
511  | 
lemma set_integral_cont_down:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
512  | 
  fixes f :: "_ \<Rightarrow> 'a :: {banach, second_countable_topology}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
513  | 
assumes [measurable]: "\<And>i. A i \<in> sets M" and A: "decseq A"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
514  | 
and int0: "set_integrable M (A 0) f"  | 
| 61969 | 515  | 
shows "(\<lambda>i::nat. LINT x:(A i)|M. f x) \<longlonglongrightarrow> LINT x:(\<Inter>i. A i)|M. f x"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
516  | 
unfolding set_lebesgue_integral_def  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
517  | 
proof (rule integral_dominated_convergence)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
518  | 
have int_A: "\<And>i. set_integrable M (A i) f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
519  | 
using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def)  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
520  | 
have "integrable M (\<lambda>c. norm (indicat_real (A 0) c *\<^sub>R f c))"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
521  | 
by (metis (no_types) int0 integrable_norm set_integrable_def)  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
522  | 
then show "integrable M (\<lambda>x. indicator (A 0) x *\<^sub>R norm (f x))"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
523  | 
by force  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
524  | 
have "set_integrable M (\<Inter>i. A i) f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
525  | 
using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def)  | 
| 69313 | 526  | 
with int_A show "(\<lambda>x. indicat_real (\<Inter>(A ` UNIV)) x *\<^sub>R f x) \<in> borel_measurable M"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
527  | 
"\<And>i. (\<lambda>x. indicat_real (A i) x *\<^sub>R f x) \<in> borel_measurable M"  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
528  | 
by (auto simp: set_integrable_def)  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
529  | 
show "\<And>i. AE x in M. norm (indicator (A i) x *\<^sub>R f x) \<le> indicator (A 0) x *\<^sub>R norm (f x)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
530  | 
using A by (auto split: split_indicator simp: decseq_def)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
531  | 
  { fix x i assume "x \<in> space M" "x \<notin> A i"
 | 
| 61969 | 532  | 
with A have "(\<lambda>i. indicator (A i) x::real) \<longlonglongrightarrow> 0 \<longleftrightarrow> (\<lambda>i. 0::real) \<longlonglongrightarrow> 0"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
533  | 
by (intro filterlim_cong refl)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
534  | 
(auto split: split_indicator simp: eventually_sequentially decseq_def intro!: exI[of _ i]) }  | 
| 61969 | 535  | 
then show "AE x in M. (\<lambda>i. indicator (A i) x *\<^sub>R f x) \<longlonglongrightarrow> indicator (\<Inter>i. A i) x *\<^sub>R f x"  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
536  | 
by (intro AE_I2 tendsto_intros) (auto split: split_indicator)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
537  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
538  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
539  | 
lemma set_integral_at_point:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
540  | 
fixes a :: real  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
541  | 
  assumes "set_integrable M {a} f"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
542  | 
  and [simp]: "{a} \<in> sets M" and "(emeasure M) {a} \<noteq> \<infinity>"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
543  | 
  shows "(LINT x:{a} | M. f x) = f a * measure M {a}"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
544  | 
proof-  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
545  | 
  have "set_lebesgue_integral M {a} f = set_lebesgue_integral M {a} (%x. f a)"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
546  | 
by (intro set_lebesgue_integral_cong) simp_all  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
547  | 
then show ?thesis using assms  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
548  | 
unfolding set_lebesgue_integral_def by simp  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
549  | 
qed  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
550  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
551  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
552  | 
abbreviation complex_integrable :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> bool" where
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
553  | 
"complex_integrable M f \<equiv> integrable M f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
554  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
555  | 
abbreviation complex_lebesgue_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> complex" ("integral\<^sup>C") where
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
556  | 
"integral\<^sup>C M f == integral\<^sup>L M f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
557  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
558  | 
syntax  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
559  | 
"_complex_lebesgue_integral" :: "pttrn \<Rightarrow> complex \<Rightarrow> 'a measure \<Rightarrow> complex"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
560  | 
 ("\<integral>\<^sup>C _. _ \<partial>_" [60,61] 110)
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
561  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
562  | 
translations  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
563  | 
"\<integral>\<^sup>Cx. f \<partial>M" == "CONST complex_lebesgue_integral M (\<lambda>x. f)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
564  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
565  | 
syntax  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
566  | 
"_ascii_complex_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
567  | 
  ("(3CLINT _|_. _)" [0,110,60] 60)
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
568  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
569  | 
translations  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
570  | 
"CLINT x|M. f" == "CONST complex_lebesgue_integral M (\<lambda>x. f)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
571  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
572  | 
lemma complex_integrable_cnj [simp]:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
573  | 
"complex_integrable M (\<lambda>x. cnj (f x)) \<longleftrightarrow> complex_integrable M f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
574  | 
proof  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
575  | 
assume "complex_integrable M (\<lambda>x. cnj (f x))"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
576  | 
then have "complex_integrable M (\<lambda>x. cnj (cnj (f x)))"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
577  | 
by (rule integrable_cnj)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
578  | 
then show "complex_integrable M f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
579  | 
by simp  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
580  | 
qed simp  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
581  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
582  | 
lemma complex_of_real_integrable_eq:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
583  | 
"complex_integrable M (\<lambda>x. complex_of_real (f x)) \<longleftrightarrow> integrable M f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
584  | 
proof  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
585  | 
assume "complex_integrable M (\<lambda>x. complex_of_real (f x))"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
586  | 
then have "integrable M (\<lambda>x. Re (complex_of_real (f x)))"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
587  | 
by (rule integrable_Re)  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
588  | 
then show "integrable M f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
589  | 
by simp  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
590  | 
qed simp  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
591  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
592  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
593  | 
abbreviation complex_set_integrable :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> bool" where
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
594  | 
"complex_set_integrable M A f \<equiv> set_integrable M A f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
595  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
596  | 
abbreviation complex_set_lebesgue_integral :: "'a measure \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> complex" where
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
597  | 
"complex_set_lebesgue_integral M A f \<equiv> set_lebesgue_integral M A f"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
598  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
599  | 
syntax  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
600  | 
"_ascii_complex_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
601  | 
("(4CLINT _:_|_. _)" [0,60,110,61] 60)
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
602  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
603  | 
translations  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
604  | 
"CLINT x:A|M. f" == "CONST complex_set_lebesgue_integral M A (\<lambda>x. f)"  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
605  | 
|
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
606  | 
lemma set_measurable_continuous_on_ivl:  | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
607  | 
  assumes "continuous_on {a..b} (f :: real \<Rightarrow> real)"
 | 
| 
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
608  | 
  shows "set_borel_measurable borel {a..b} f"
 | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
609  | 
unfolding set_borel_measurable_def  | 
| 
66164
 
2d79288b042c
New theorems and much tidying up of the old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
64911 
diff
changeset
 | 
610  | 
by (rule borel_measurable_continuous_on_indicator[OF _ assms]) simp  | 
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
611  | 
|
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
612  | 
|
| 64911 | 613  | 
text\<open>This notation is from Sébastien Gouëzel: His use is not directly in line with the  | 
614  | 
notations in this file, they are more in line with sum, and more readable he thinks.\<close>  | 
|
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
615  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
616  | 
abbreviation "set_nn_integral M A f \<equiv> nn_integral M (\<lambda>x. f x * indicator A x)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
617  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
618  | 
syntax  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
619  | 
"_set_nn_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
620  | 
("(\<integral>\<^sup>+((_)\<in>(_)./ _)/\<partial>_)" [0,60,110,61] 60)
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
621  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
622  | 
"_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
623  | 
("(\<integral>((_)\<in>(_)./ _)/\<partial>_)" [0,60,110,61] 60)
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
624  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
625  | 
translations  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
626  | 
"\<integral>\<^sup>+x \<in> A. f \<partial>M" == "CONST set_nn_integral M A (\<lambda>x. f)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
627  | 
"\<integral>x \<in> A. f \<partial>M" == "CONST set_lebesgue_integral M A (\<lambda>x. f)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
628  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
629  | 
lemma nn_integral_disjoint_pair:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
630  | 
assumes [measurable]: "f \<in> borel_measurable M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
631  | 
"B \<in> sets M" "C \<in> sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
632  | 
          "B \<inter> C = {}"
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
633  | 
shows "(\<integral>\<^sup>+x \<in> B \<union> C. f x \<partial>M) = (\<integral>\<^sup>+x \<in> B. f x \<partial>M) + (\<integral>\<^sup>+x \<in> C. f x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
634  | 
proof -  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
635  | 
have mes: "\<And>D. D \<in> sets M \<Longrightarrow> (\<lambda>x. f x * indicator D x) \<in> borel_measurable M" by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
636  | 
have pos: "\<And>D. AE x in M. f x * indicator D x \<ge> 0" using assms(2) by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
637  | 
have "\<And>x. f x * indicator (B \<union> C) x = f x * indicator B x + f x * indicator C x" using assms(4)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
638  | 
by (auto split: split_indicator)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
639  | 
then have "(\<integral>\<^sup>+x. f x * indicator (B \<union> C) x \<partial>M) = (\<integral>\<^sup>+x. f x * indicator B x + f x * indicator C x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
640  | 
by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
641  | 
also have "... = (\<integral>\<^sup>+x. f x * indicator B x \<partial>M) + (\<integral>\<^sup>+x. f x * indicator C x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
642  | 
by (rule nn_integral_add) (auto simp add: assms mes pos)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
643  | 
finally show ?thesis by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
644  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
645  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
646  | 
lemma nn_integral_disjoint_pair_countspace:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
647  | 
  assumes "B \<inter> C = {}"
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
648  | 
shows "(\<integral>\<^sup>+x \<in> B \<union> C. f x \<partial>count_space UNIV) = (\<integral>\<^sup>+x \<in> B. f x \<partial>count_space UNIV) + (\<integral>\<^sup>+x \<in> C. f x \<partial>count_space UNIV)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
649  | 
by (rule nn_integral_disjoint_pair) (simp_all add: assms)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
650  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
651  | 
lemma nn_integral_null_delta:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
652  | 
assumes "A \<in> sets M" "B \<in> sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
653  | 
"(A - B) \<union> (B - A) \<in> null_sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
654  | 
shows "(\<integral>\<^sup>+x \<in> A. f x \<partial>M) = (\<integral>\<^sup>+x \<in> B. f x \<partial>M)"  | 
| 73536 | 655  | 
proof (rule nn_integral_cong_AE)  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
656  | 
have *: "AE a in M. a \<notin> (A - B) \<union> (B - A)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
657  | 
using assms(3) AE_not_in by blast  | 
| 73536 | 658  | 
then show \<open>AE x in M. f x * indicator A x = f x * indicator B x\<close>  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
659  | 
by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
660  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
661  | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
662  | 
proposition nn_integral_disjoint_family:  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
663  | 
assumes [measurable]: "f \<in> borel_measurable M" "\<And>(n::nat). B n \<in> sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
664  | 
and "disjoint_family B"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
665  | 
shows "(\<integral>\<^sup>+x \<in> (\<Union>n. B n). f x \<partial>M) = (\<Sum>n. (\<integral>\<^sup>+x \<in> B n. f x \<partial>M))"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
666  | 
proof -  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
667  | 
have "(\<integral>\<^sup>+ x. (\<Sum>n. f x * indicator (B n) x) \<partial>M) = (\<Sum>n. (\<integral>\<^sup>+ x. f x * indicator (B n) x \<partial>M))"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
668  | 
by (rule nn_integral_suminf) simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
669  | 
moreover have "(\<Sum>n. f x * indicator (B n) x) = f x * indicator (\<Union>n. B n) x" for x  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
670  | 
proof (cases)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
671  | 
assume "x \<in> (\<Union>n. B n)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
672  | 
then obtain n where "x \<in> B n" by blast  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
673  | 
    have a: "finite {n}" by simp
 | 
| 64911 | 674  | 
have "\<And>i. i \<noteq> n \<Longrightarrow> x \<notin> B i" using \<open>x \<in> B n\<close> assms(3) disjoint_family_on_def  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
675  | 
by (metis IntI UNIV_I empty_iff)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
676  | 
    then have "\<And>i. i \<notin> {n} \<Longrightarrow> indicator (B i) x = (0::ennreal)" using indicator_def by simp
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
677  | 
    then have b: "\<And>i. i \<notin> {n} \<Longrightarrow> f x * indicator (B i) x = (0::ennreal)" by simp
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
678  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
679  | 
define h where "h = (\<lambda>i. f x * indicator (B i) x)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
680  | 
    then have "\<And>i. i \<notin> {n} \<Longrightarrow> h i = 0" using b by simp
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
681  | 
    then have "(\<Sum>i. h i) = (\<Sum>i\<in>{n}. h i)"
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
682  | 
by (metis sums_unique[OF sums_finite[OF a]])  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
683  | 
then have "(\<Sum>i. h i) = h n" by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
684  | 
then have "(\<Sum>n. f x * indicator (B n) x) = f x * indicator (B n) x" using h_def by simp  | 
| 64911 | 685  | 
then have "(\<Sum>n. f x * indicator (B n) x) = f x" using \<open>x \<in> B n\<close> indicator_def by simp  | 
686  | 
then show ?thesis using \<open>x \<in> (\<Union>n. B n)\<close> by auto  | 
|
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
687  | 
next  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
688  | 
assume "x \<notin> (\<Union>n. B n)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
689  | 
then have "\<And>n. f x * indicator (B n) x = 0" by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
690  | 
have "(\<Sum>n. f x * indicator (B n) x) = 0"  | 
| 64911 | 691  | 
by (simp add: \<open>\<And>n. f x * indicator (B n) x = 0\<close>)  | 
692  | 
then show ?thesis using \<open>x \<notin> (\<Union>n. B n)\<close> by auto  | 
|
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
693  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
694  | 
ultimately show ?thesis by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
695  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
696  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
697  | 
lemma nn_set_integral_add:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
698  | 
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
699  | 
"A \<in> sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
700  | 
shows "(\<integral>\<^sup>+x \<in> A. (f x + g x) \<partial>M) = (\<integral>\<^sup>+x \<in> A. f x \<partial>M) + (\<integral>\<^sup>+x \<in> A. g x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
701  | 
proof -  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
702  | 
have "(\<integral>\<^sup>+x \<in> A. (f x + g x) \<partial>M) = (\<integral>\<^sup>+x. (f x * indicator A x + g x * indicator A x) \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
703  | 
by (auto simp add: indicator_def intro!: nn_integral_cong)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
704  | 
also have "... = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) + (\<integral>\<^sup>+x. g x * indicator A x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
705  | 
apply (rule nn_integral_add) using assms(1) assms(2) by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
706  | 
finally show ?thesis by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
707  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
708  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
709  | 
lemma nn_set_integral_cong:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
710  | 
assumes "AE x in M. f x = g x"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
711  | 
shows "(\<integral>\<^sup>+x \<in> A. f x \<partial>M) = (\<integral>\<^sup>+x \<in> A. g x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
712  | 
apply (rule nn_integral_cong_AE) using assms(1) by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
713  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
714  | 
lemma nn_set_integral_set_mono:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
715  | 
"A \<subseteq> B \<Longrightarrow> (\<integral>\<^sup>+ x \<in> A. f x \<partial>M) \<le> (\<integral>\<^sup>+ x \<in> B. f x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
716  | 
by (auto intro!: nn_integral_mono split: split_indicator)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
717  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
718  | 
lemma nn_set_integral_mono:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
719  | 
assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
720  | 
"A \<in> sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
721  | 
and "AE x\<in>A in M. f x \<le> g x"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
722  | 
shows "(\<integral>\<^sup>+x \<in> A. f x \<partial>M) \<le> (\<integral>\<^sup>+x \<in> A. g x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
723  | 
by (auto intro!: nn_integral_mono_AE split: split_indicator simp: assms)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
724  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
725  | 
lemma nn_set_integral_space [simp]:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
726  | 
shows "(\<integral>\<^sup>+ x \<in> space M. f x \<partial>M) = (\<integral>\<^sup>+x. f x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
727  | 
by (metis (mono_tags, lifting) indicator_simps(1) mult.right_neutral nn_integral_cong)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
728  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
729  | 
lemma nn_integral_count_compose_inj:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
730  | 
assumes "inj_on g A"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
731  | 
shows "(\<integral>\<^sup>+x \<in> A. f (g x) \<partial>count_space UNIV) = (\<integral>\<^sup>+y \<in> g`A. f y \<partial>count_space UNIV)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
732  | 
proof -  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
733  | 
have "(\<integral>\<^sup>+x \<in> A. f (g x) \<partial>count_space UNIV) = (\<integral>\<^sup>+x. f (g x) \<partial>count_space A)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
734  | 
by (auto simp add: nn_integral_count_space_indicator[symmetric])  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
735  | 
also have "... = (\<integral>\<^sup>+y. f y \<partial>count_space (g`A))"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
736  | 
by (simp add: assms nn_integral_bij_count_space inj_on_imp_bij_betw)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
737  | 
also have "... = (\<integral>\<^sup>+y \<in> g`A. f y \<partial>count_space UNIV)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
738  | 
by (auto simp add: nn_integral_count_space_indicator[symmetric])  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
739  | 
finally show ?thesis by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
740  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
741  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
742  | 
lemma nn_integral_count_compose_bij:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
743  | 
assumes "bij_betw g A B"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
744  | 
shows "(\<integral>\<^sup>+x \<in> A. f (g x) \<partial>count_space UNIV) = (\<integral>\<^sup>+y \<in> B. f y \<partial>count_space UNIV)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
745  | 
proof -  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
746  | 
have "inj_on g A" using assms bij_betw_def by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
747  | 
then have "(\<integral>\<^sup>+x \<in> A. f (g x) \<partial>count_space UNIV) = (\<integral>\<^sup>+y \<in> g`A. f y \<partial>count_space UNIV)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
748  | 
by (rule nn_integral_count_compose_inj)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
749  | 
then show ?thesis using assms by (simp add: bij_betw_def)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
750  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
751  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
752  | 
lemma set_integral_null_delta:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
753  | 
  fixes f::"_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
754  | 
assumes [measurable]: "integrable M f" "A \<in> sets M" "B \<in> sets M"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
755  | 
and null: "(A - B) \<union> (B - A) \<in> null_sets M"  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
756  | 
shows "(\<integral>x \<in> A. f x \<partial>M) = (\<integral>x \<in> B. f x \<partial>M)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
757  | 
proof (rule set_integral_cong_set)  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
758  | 
have *: "AE a in M. a \<notin> (A - B) \<union> (B - A)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
759  | 
using null AE_not_in by blast  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
760  | 
then show "AE x in M. (x \<in> B) = (x \<in> A)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
761  | 
by auto  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
762  | 
qed (simp_all add: set_borel_measurable_def)  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
763  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
764  | 
lemma set_integral_space:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
765  | 
assumes "integrable M f"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
766  | 
shows "(\<integral>x \<in> space M. f x \<partial>M) = (\<integral>x. f x \<partial>M)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
767  | 
by (metis (no_types, lifting) indicator_simps(1) integral_cong scaleR_one set_lebesgue_integral_def)  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
768  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
769  | 
lemma null_if_pos_func_has_zero_nn_int:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
770  | 
fixes f::"'a \<Rightarrow> ennreal"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
771  | 
assumes [measurable]: "f \<in> borel_measurable M" "A \<in> sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
772  | 
and "AE x\<in>A in M. f x > 0" "(\<integral>\<^sup>+x\<in>A. f x \<partial>M) = 0"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
773  | 
shows "A \<in> null_sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
774  | 
proof -  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
775  | 
have "AE x in M. f x * indicator A x = 0"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
776  | 
by (subst nn_integral_0_iff_AE[symmetric], auto simp add: assms(4))  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
777  | 
then have "AE x\<in>A in M. False" using assms(3) by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
778  | 
then show "A \<in> null_sets M" using assms(2) by (simp add: AE_iff_null_sets)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
779  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
780  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
781  | 
lemma null_if_pos_func_has_zero_int:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
782  | 
assumes [measurable]: "integrable M f" "A \<in> sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
783  | 
and "AE x\<in>A in M. f x > 0" "(\<integral>x\<in>A. f x \<partial>M) = (0::real)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
784  | 
shows "A \<in> null_sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
785  | 
proof -  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
786  | 
have "AE x in M. indicator A x * f x = 0"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
787  | 
apply (subst integral_nonneg_eq_0_iff_AE[symmetric])  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
788  | 
using assms integrable_mult_indicator[OF \<open>A \<in> sets M\<close> assms(1)]  | 
| 
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
789  | 
by (auto simp: set_lebesgue_integral_def)  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
790  | 
then have "AE x\<in>A in M. f x = 0" by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
791  | 
then have "AE x\<in>A in M. False" using assms(3) by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
792  | 
then show "A \<in> null_sets M" using assms(2) by (simp add: AE_iff_null_sets)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
793  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
794  | 
|
| 64911 | 795  | 
text\<open>The next lemma is a variant of \<open>density_unique\<close>. Note that it uses the notation  | 
796  | 
for nonnegative set integrals introduced earlier.\<close>  | 
|
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
797  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
798  | 
lemma (in sigma_finite_measure) density_unique2:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
799  | 
assumes [measurable]: "f \<in> borel_measurable M" "f' \<in> borel_measurable M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
800  | 
assumes density_eq: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>\<^sup>+ x \<in> A. f x \<partial>M) = (\<integral>\<^sup>+ x \<in> A. f' x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
801  | 
shows "AE x in M. f x = f' x"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
802  | 
proof (rule density_unique)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
803  | 
show "density M f = density M f'"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
804  | 
by (intro measure_eqI) (auto simp: emeasure_density intro!: density_eq)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
805  | 
qed (auto simp add: assms)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
806  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
807  | 
|
| 64911 | 808  | 
text \<open>The next lemma implies the same statement for Banach-space valued functions  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
809  | 
using Hahn-Banach theorem and linear forms. Since they are not yet easily available, I  | 
| 64911 | 810  | 
only formulate it for real-valued functions.\<close>  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
811  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
812  | 
lemma density_unique_real:  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
813  | 
fixes f f'::"_ \<Rightarrow> real"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
814  | 
assumes M[measurable]: "integrable M f" "integrable M f'"  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
815  | 
assumes density_eq: "\<And>A. A \<in> sets M \<Longrightarrow> (\<integral>x \<in> A. f x \<partial>M) = (\<integral>x \<in> A. f' x \<partial>M)"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
816  | 
shows "AE x in M. f x = f' x"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
817  | 
proof -  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
818  | 
  define A where "A = {x \<in> space M. f x < f' x}"
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
819  | 
then have [measurable]: "A \<in> sets M" by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
820  | 
have "(\<integral>x \<in> A. (f' x - f x) \<partial>M) = (\<integral>x \<in> A. f' x \<partial>M) - (\<integral>x \<in> A. f x \<partial>M)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
821  | 
using \<open>A \<in> sets M\<close> M integrable_mult_indicator set_integrable_def by blast  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
822  | 
then have "(\<integral>x \<in> A. (f' x - f x) \<partial>M) = 0" using assms(3) by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
823  | 
then have "A \<in> null_sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
824  | 
using A_def null_if_pos_func_has_zero_int[where ?f = "\<lambda>x. f' x - f x" and ?A = A] assms by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
825  | 
then have "AE x in M. x \<notin> A" by (simp add: AE_not_in)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
826  | 
then have *: "AE x in M. f' x \<le> f x" unfolding A_def by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
827  | 
|
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
828  | 
  define B where "B = {x \<in> space M. f' x < f x}"
 | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
829  | 
then have [measurable]: "B \<in> sets M" by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
830  | 
have "(\<integral>x \<in> B. (f x - f' x) \<partial>M) = (\<integral>x \<in> B. f x \<partial>M) - (\<integral>x \<in> B. f' x \<partial>M)"  | 
| 
67974
 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 
paulson <lp15@cam.ac.uk> 
parents: 
67399 
diff
changeset
 | 
831  | 
using \<open>B \<in> sets M\<close> M integrable_mult_indicator set_integrable_def by blast  | 
| 
64283
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
832  | 
then have "(\<integral>x \<in> B. (f x - f' x) \<partial>M) = 0" using assms(3) by simp  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
833  | 
then have "B \<in> null_sets M"  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
834  | 
using B_def null_if_pos_func_has_zero_int[where ?f = "\<lambda>x. f x - f' x" and ?A = B] assms by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
835  | 
then have "AE x in M. x \<notin> B" by (simp add: AE_not_in)  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
836  | 
then have "AE x in M. f' x \<ge> f x" unfolding B_def by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
837  | 
then show ?thesis using * by auto  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
838  | 
qed  | 
| 
 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 
hoelzl 
parents: 
63958 
diff
changeset
 | 
839  | 
|
| 69566 | 840  | 
text \<open>The next lemma shows that \<open>L\<^sup>1\<close> convergence of a sequence of functions follows from almost  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
841  | 
everywhere convergence and the weaker condition of the convergence of the integrated norms (or even  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
842  | 
just the nontrivial inequality about them). Useful in a lot of contexts! This statement (or its  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
843  | 
variations) are known as Scheffe lemma.  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
844  | 
|
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
845  | 
The formalization is more painful as one should jump back and forth between reals and ereals and justify  | 
| 64911 | 846  | 
all the time positivity or integrability (thankfully, measurability is handled more or less automatically).\<close>  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
847  | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
848  | 
proposition Scheffe_lemma1:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
849  | 
assumes "\<And>n. integrable M (F n)" "integrable M f"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
850  | 
"AE x in M. (\<lambda>n. F n x) \<longlonglongrightarrow> f x"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
851  | 
"limsup (\<lambda>n. \<integral>\<^sup>+ x. norm(F n x) \<partial>M) \<le> (\<integral>\<^sup>+ x. norm(f x) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
852  | 
shows "(\<lambda>n. \<integral>\<^sup>+ x. norm(F n x - f x) \<partial>M) \<longlonglongrightarrow> 0"  | 
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
853  | 
proof -  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
854  | 
have [measurable]: "\<And>n. F n \<in> borel_measurable M" "f \<in> borel_measurable M"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
855  | 
using assms(1) assms(2) by simp_all  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
856  | 
define G where "G = (\<lambda>n x. norm(f x) + norm(F n x) - norm(F n x - f x))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
857  | 
have [measurable]: "\<And>n. G n \<in> borel_measurable M" unfolding G_def by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
858  | 
have G_pos[simp]: "\<And>n x. G n x \<ge> 0"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
859  | 
unfolding G_def by (metis ge_iff_diff_ge_0 norm_minus_commute norm_triangle_ineq4)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
860  | 
have finint: "(\<integral>\<^sup>+ x. norm(f x) \<partial>M) \<noteq> \<infinity>"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
861  | 
using has_bochner_integral_implies_finite_norm[OF has_bochner_integral_integrable[OF \<open>integrable M f\<close>]]  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
862  | 
by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
863  | 
then have fin2: "2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M) \<noteq> \<infinity>"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
864  | 
by (auto simp: ennreal_mult_eq_top_iff)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
865  | 
|
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
866  | 
  {
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
867  | 
fix x assume *: "(\<lambda>n. F n x) \<longlonglongrightarrow> f x"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
868  | 
then have "(\<lambda>n. norm(F n x)) \<longlonglongrightarrow> norm(f x)" using tendsto_norm by blast  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
869  | 
moreover have "(\<lambda>n. norm(F n x - f x)) \<longlonglongrightarrow> 0" using * Lim_null tendsto_norm_zero_iff by fastforce  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
870  | 
ultimately have a: "(\<lambda>n. norm(F n x) - norm(F n x - f x)) \<longlonglongrightarrow> norm(f x)" using tendsto_diff by fastforce  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
871  | 
have "(\<lambda>n. norm(f x) + (norm(F n x) - norm(F n x - f x))) \<longlonglongrightarrow> norm(f x) + norm(f x)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
872  | 
by (rule tendsto_add) (auto simp add: a)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
873  | 
moreover have "\<And>n. G n x = norm(f x) + (norm(F n x) - norm(F n x - f x))" unfolding G_def by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
874  | 
ultimately have "(\<lambda>n. G n x) \<longlonglongrightarrow> 2 * norm(f x)" by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
875  | 
then have "(\<lambda>n. ennreal(G n x)) \<longlonglongrightarrow> ennreal(2 * norm(f x))" by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
876  | 
then have "liminf (\<lambda>n. ennreal(G n x)) = ennreal(2 * norm(f x))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
877  | 
using sequentially_bot tendsto_iff_Liminf_eq_Limsup by blast  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
878  | 
}  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
879  | 
then have "AE x in M. liminf (\<lambda>n. ennreal(G n x)) = ennreal(2 * norm(f x))" using assms(3) by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
880  | 
then have "(\<integral>\<^sup>+ x. liminf (\<lambda>n. ennreal (G n x)) \<partial>M) = (\<integral>\<^sup>+ x. 2 * ennreal(norm(f x)) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
881  | 
by (simp add: nn_integral_cong_AE ennreal_mult)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
882  | 
also have "... = 2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M)" by (rule nn_integral_cmult) auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
883  | 
finally have int_liminf: "(\<integral>\<^sup>+ x. liminf (\<lambda>n. ennreal (G n x)) \<partial>M) = 2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
884  | 
by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
885  | 
|
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
886  | 
have "(\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) = (\<integral>\<^sup>+x. norm(f x) \<partial>M) + (\<integral>\<^sup>+x. norm(F n x) \<partial>M)" for n  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
887  | 
by (rule nn_integral_add) (auto simp add: assms)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
888  | 
then have "limsup (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M)) =  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
889  | 
limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(f x) \<partial>M) + (\<integral>\<^sup>+x. norm(F n x) \<partial>M))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
890  | 
by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
891  | 
also have "... = (\<integral>\<^sup>+x. norm(f x) \<partial>M) + limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x) \<partial>M))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
892  | 
by (rule Limsup_const_add, auto simp add: finint)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
893  | 
also have "... \<le> (\<integral>\<^sup>+x. norm(f x) \<partial>M) + (\<integral>\<^sup>+x. norm(f x) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
894  | 
using assms(4) by (simp add: add_left_mono)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
895  | 
also have "... = 2 * (\<integral>\<^sup>+x. norm(f x) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
896  | 
unfolding one_add_one[symmetric] distrib_right by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
897  | 
ultimately have a: "limsup (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M)) \<le>  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
898  | 
2 * (\<integral>\<^sup>+x. norm(f x) \<partial>M)" by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
899  | 
|
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
900  | 
have le: "ennreal (norm (F n x - f x)) \<le> ennreal (norm (f x)) + ennreal (norm (F n x))" for n x  | 
| 68403 | 901  | 
by (simp add: norm_minus_commute norm_triangle_ineq4 ennreal_minus flip: ennreal_plus)  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
902  | 
then have le2: "(\<integral>\<^sup>+ x. ennreal (norm (F n x - f x)) \<partial>M) \<le> (\<integral>\<^sup>+ x. ennreal (norm (f x)) + ennreal (norm (F n x)) \<partial>M)" for n  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
903  | 
by (rule nn_integral_mono)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
904  | 
|
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
905  | 
have "2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M) = (\<integral>\<^sup>+ x. liminf (\<lambda>n. ennreal (G n x)) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
906  | 
by (simp add: int_liminf)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
907  | 
also have "\<dots> \<le> liminf (\<lambda>n. (\<integral>\<^sup>+x. G n x \<partial>M))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
908  | 
by (rule nn_integral_liminf) auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
909  | 
also have "liminf (\<lambda>n. (\<integral>\<^sup>+x. G n x \<partial>M)) =  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
910  | 
liminf (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
911  | 
proof (intro arg_cong[where f=liminf] ext)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
912  | 
fix n  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
913  | 
have "\<And>x. ennreal(G n x) = ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x))"  | 
| 68403 | 914  | 
unfolding G_def by (simp add: ennreal_minus flip: ennreal_plus)  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
915  | 
moreover have "(\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x)) \<partial>M)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
916  | 
= (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
917  | 
proof (rule nn_integral_diff)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
918  | 
from le show "AE x in M. ennreal (norm (F n x - f x)) \<le> ennreal (norm (f x)) + ennreal (norm (F n x))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
919  | 
by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
920  | 
from le2 have "(\<integral>\<^sup>+x. ennreal (norm (F n x - f x)) \<partial>M) < \<infinity>" using assms(1) assms(2)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
921  | 
by (metis has_bochner_integral_implies_finite_norm integrable.simps Bochner_Integration.integrable_diff)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
922  | 
then show "(\<integral>\<^sup>+x. ennreal (norm (F n x - f x)) \<partial>M) \<noteq> \<infinity>" by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
923  | 
qed (auto simp add: assms)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
924  | 
ultimately show "(\<integral>\<^sup>+x. G n x \<partial>M) = (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
925  | 
by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
926  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
927  | 
finally have "2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M) + limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)) \<le>  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
928  | 
liminf (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)) +  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
929  | 
limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
930  | 
by (intro add_mono) auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
931  | 
also have "\<dots> \<le> (limsup (\<lambda>n. \<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M) - limsup (\<lambda>n. \<integral>\<^sup>+x. norm (F n x - f x) \<partial>M)) +  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
932  | 
limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
933  | 
by (intro add_mono liminf_minus_ennreal le2) auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
934  | 
also have "\<dots> = limsup (\<lambda>n. (\<integral>\<^sup>+x. ennreal(norm(f x)) + ennreal(norm(F n x)) \<partial>M))"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
935  | 
by (intro diff_add_cancel_ennreal Limsup_mono always_eventually allI le2)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
936  | 
also have "\<dots> \<le> 2 * (\<integral>\<^sup>+x. norm(f x) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
937  | 
by fact  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
938  | 
finally have "limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x - f x) \<partial>M)) = 0"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
939  | 
using fin2 by simp  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
940  | 
then show ?thesis  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
941  | 
by (rule tendsto_0_if_Limsup_eq_0_ennreal)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
942  | 
qed  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
943  | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
944  | 
proposition Scheffe_lemma2:  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
945  | 
  fixes F::"nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
946  | 
assumes "\<And> n::nat. F n \<in> borel_measurable M" "integrable M f"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
947  | 
"AE x in M. (\<lambda>n. F n x) \<longlonglongrightarrow> f x"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
948  | 
"\<And>n. (\<integral>\<^sup>+ x. norm(F n x) \<partial>M) \<le> (\<integral>\<^sup>+ x. norm(f x) \<partial>M)"  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
949  | 
shows "(\<lambda>n. \<integral>\<^sup>+ x. norm(F n x - f x) \<partial>M) \<longlonglongrightarrow> 0"  | 
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
950  | 
proof (rule Scheffe_lemma1)  | 
| 
64284
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
951  | 
fix n::nat  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
952  | 
have "(\<integral>\<^sup>+ x. norm(f x) \<partial>M) < \<infinity>" using assms(2) by (metis has_bochner_integral_implies_finite_norm integrable.cases)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
953  | 
then have "(\<integral>\<^sup>+ x. norm(F n x) \<partial>M) < \<infinity>" using assms(4)[of n] by auto  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
954  | 
then show "integrable M (F n)" by (subst integrable_iff_bounded, simp add: assms(1)[of n])  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
955  | 
qed (auto simp add: assms Limsup_bounded)  | 
| 
 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 
hoelzl 
parents: 
64283 
diff
changeset
 | 
956  | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
957  | 
lemma tendsto_set_lebesgue_integral_at_right:  | 
| 68721 | 958  | 
  fixes a b :: real and f :: "real \<Rightarrow> 'a :: {banach,second_countable_topology}"
 | 
959  | 
  assumes "a < b" and sets: "\<And>a'. a' \<in> {a<..b} \<Longrightarrow> {a'..b} \<in> sets M"
 | 
|
960  | 
      and "set_integrable M {a<..b} f"
 | 
|
961  | 
  shows   "((\<lambda>a'. set_lebesgue_integral M {a'..b} f) \<longlongrightarrow>
 | 
|
962  | 
             set_lebesgue_integral M {a<..b} f) (at_right a)"
 | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
963  | 
proof (rule tendsto_at_right_sequentially[OF assms(1)], goal_cases)  | 
| 68721 | 964  | 
case (1 S)  | 
965  | 
  have eq: "(\<Union>n. {S n..b}) = {a<..b}"
 | 
|
966  | 
proof safe  | 
|
967  | 
    fix x n assume "x \<in> {S n..b}"
 | 
|
968  | 
    with 1(1,2)[of n] show "x \<in> {a<..b}" by auto
 | 
|
969  | 
next  | 
|
970  | 
    fix x assume "x \<in> {a<..b}"
 | 
|
971  | 
    with order_tendstoD[OF \<open>S \<longlonglongrightarrow> a\<close>, of x] show "x \<in> (\<Union>n. {S n..b})"
 | 
|
972  | 
by (force simp: eventually_at_top_linorder dest: less_imp_le)  | 
|
973  | 
qed  | 
|
974  | 
  have "(\<lambda>n. set_lebesgue_integral M {S n..b} f) \<longlonglongrightarrow> set_lebesgue_integral M (\<Union>n. {S n..b}) f"
 | 
|
975  | 
by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le)  | 
|
976  | 
with eq show ?case by simp  | 
|
977  | 
qed  | 
|
978  | 
||
979  | 
||
980  | 
text \<open>  | 
|
981  | 
The next lemmas relate convergence of integrals over an interval to  | 
|
982  | 
improper integrals.  | 
|
983  | 
\<close>  | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
984  | 
lemma tendsto_set_lebesgue_integral_at_left:  | 
| 68721 | 985  | 
  fixes a b :: real and f :: "real \<Rightarrow> 'a :: {banach,second_countable_topology}"
 | 
986  | 
  assumes "a < b" and sets: "\<And>b'. b' \<in> {a..<b} \<Longrightarrow> {a..b'} \<in> sets M"
 | 
|
987  | 
      and "set_integrable M {a..<b} f"
 | 
|
988  | 
  shows   "((\<lambda>b'. set_lebesgue_integral M {a..b'} f) \<longlongrightarrow>
 | 
|
989  | 
             set_lebesgue_integral M {a..<b} f) (at_left b)"
 | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
990  | 
proof (rule tendsto_at_left_sequentially[OF assms(1)], goal_cases)  | 
| 68721 | 991  | 
case (1 S)  | 
992  | 
  have eq: "(\<Union>n. {a..S n}) = {a..<b}"
 | 
|
993  | 
proof safe  | 
|
994  | 
    fix x n assume "x \<in> {a..S n}"
 | 
|
995  | 
    with 1(1,2)[of n] show "x \<in> {a..<b}" by auto
 | 
|
996  | 
next  | 
|
997  | 
    fix x assume "x \<in> {a..<b}"
 | 
|
998  | 
    with order_tendstoD[OF \<open>S \<longlonglongrightarrow> b\<close>, of x] show "x \<in> (\<Union>n. {a..S n})"
 | 
|
999  | 
by (force simp: eventually_at_top_linorder dest: less_imp_le)  | 
|
1000  | 
qed  | 
|
1001  | 
  have "(\<lambda>n. set_lebesgue_integral M {a..S n} f) \<longlonglongrightarrow> set_lebesgue_integral M (\<Union>n. {a..S n}) f"
 | 
|
1002  | 
by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le)  | 
|
1003  | 
with eq show ?case by simp  | 
|
1004  | 
qed  | 
|
1005  | 
||
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
1006  | 
proposition tendsto_set_lebesgue_integral_at_top:  | 
| 68721 | 1007  | 
  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
 | 
1008  | 
  assumes sets: "\<And>b. b \<ge> a \<Longrightarrow> {a..b} \<in> sets M"
 | 
|
1009  | 
      and int: "set_integrable M {a..} f"
 | 
|
1010  | 
  shows "((\<lambda>b. set_lebesgue_integral M {a..b} f) \<longlongrightarrow> set_lebesgue_integral M {a..} f) at_top"
 | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
1011  | 
proof (rule tendsto_at_topI_sequentially)  | 
| 68721 | 1012  | 
fix X :: "nat \<Rightarrow> real" assume "filterlim X at_top sequentially"  | 
1013  | 
  show "(\<lambda>n. set_lebesgue_integral M {a..X n} f) \<longlonglongrightarrow> set_lebesgue_integral M {a..} f"
 | 
|
1014  | 
unfolding set_lebesgue_integral_def  | 
|
1015  | 
proof (rule integral_dominated_convergence)  | 
|
1016  | 
    show "integrable M (\<lambda>x. indicat_real {a..} x *\<^sub>R norm (f x))"
 | 
|
1017  | 
using integrable_norm[OF int[unfolded set_integrable_def]] by simp  | 
|
1018  | 
    show "AE x in M. (\<lambda>n. indicator {a..X n} x *\<^sub>R f x) \<longlonglongrightarrow> indicat_real {a..} x *\<^sub>R f x"
 | 
|
1019  | 
proof  | 
|
1020  | 
fix x  | 
|
1021  | 
from \<open>filterlim X at_top sequentially\<close>  | 
|
1022  | 
have "eventually (\<lambda>n. x \<le> X n) sequentially"  | 
|
1023  | 
unfolding filterlim_at_top_ge[where c=x] by auto  | 
|
1024  | 
      then show "(\<lambda>n. indicator {a..X n} x *\<^sub>R f x) \<longlonglongrightarrow> indicat_real {a..} x *\<^sub>R f x"
 | 
|
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1025  | 
by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono)  | 
| 68721 | 1026  | 
qed  | 
1027  | 
    fix n show "AE x in M. norm (indicator {a..X n} x *\<^sub>R f x) \<le> 
 | 
|
1028  | 
                             indicator {a..} x *\<^sub>R norm (f x)"
 | 
|
1029  | 
by (auto split: split_indicator)  | 
|
1030  | 
next  | 
|
1031  | 
    from int show "(\<lambda>x. indicat_real {a..} x *\<^sub>R f x) \<in> borel_measurable M"
 | 
|
1032  | 
by (simp add: set_integrable_def)  | 
|
1033  | 
next  | 
|
1034  | 
fix n :: nat  | 
|
1035  | 
    from sets have "{a..X n} \<in> sets M" by (cases "X n \<ge> a") auto
 | 
|
1036  | 
    with int have "set_integrable M {a..X n} f"
 | 
|
1037  | 
by (rule set_integrable_subset) auto  | 
|
1038  | 
    thus "(\<lambda>x. indicat_real {a..X n} x *\<^sub>R f x) \<in> borel_measurable M"
 | 
|
1039  | 
by (simp add: set_integrable_def)  | 
|
1040  | 
qed  | 
|
1041  | 
qed  | 
|
1042  | 
||
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
1043  | 
proposition tendsto_set_lebesgue_integral_at_bot:  | 
| 68721 | 1044  | 
  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
 | 
1045  | 
  assumes sets: "\<And>a. a \<le> b \<Longrightarrow> {a..b} \<in> sets M"
 | 
|
1046  | 
      and int: "set_integrable M {..b} f"
 | 
|
1047  | 
    shows "((\<lambda>a. set_lebesgue_integral M {a..b} f) \<longlongrightarrow> set_lebesgue_integral M {..b} f) at_bot"
 | 
|
| 
69737
 
ec3cc98c38db
tagged 4 theories
 
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> 
parents: 
69566 
diff
changeset
 | 
1048  | 
proof (rule tendsto_at_botI_sequentially)  | 
| 68721 | 1049  | 
fix X :: "nat \<Rightarrow> real" assume "filterlim X at_bot sequentially"  | 
1050  | 
  show "(\<lambda>n. set_lebesgue_integral M {X n..b} f) \<longlonglongrightarrow> set_lebesgue_integral M {..b} f"
 | 
|
1051  | 
unfolding set_lebesgue_integral_def  | 
|
1052  | 
proof (rule integral_dominated_convergence)  | 
|
1053  | 
    show "integrable M (\<lambda>x. indicat_real {..b} x *\<^sub>R norm (f x))"
 | 
|
1054  | 
using integrable_norm[OF int[unfolded set_integrable_def]] by simp  | 
|
1055  | 
    show "AE x in M. (\<lambda>n. indicator {X n..b} x *\<^sub>R f x) \<longlonglongrightarrow> indicat_real {..b} x *\<^sub>R f x"
 | 
|
1056  | 
proof  | 
|
1057  | 
fix x  | 
|
1058  | 
from \<open>filterlim X at_bot sequentially\<close>  | 
|
1059  | 
have "eventually (\<lambda>n. x \<ge> X n) sequentially"  | 
|
1060  | 
unfolding filterlim_at_bot_le[where c=x] by auto  | 
|
1061  | 
      then show "(\<lambda>n. indicator {X n..b} x *\<^sub>R f x) \<longlonglongrightarrow> indicat_real {..b} x *\<^sub>R f x"
 | 
|
| 
70365
 
4df0628e8545
a few new lemmas and a bit of tidying
 
paulson <lp15@cam.ac.uk> 
parents: 
70136 
diff
changeset
 | 
1062  | 
by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono)  | 
| 68721 | 1063  | 
qed  | 
1064  | 
    fix n show "AE x in M. norm (indicator {X n..b} x *\<^sub>R f x) \<le> 
 | 
|
1065  | 
                             indicator {..b} x *\<^sub>R norm (f x)"
 | 
|
1066  | 
by (auto split: split_indicator)  | 
|
1067  | 
next  | 
|
1068  | 
    from int show "(\<lambda>x. indicat_real {..b} x *\<^sub>R f x) \<in> borel_measurable M"
 | 
|
1069  | 
by (simp add: set_integrable_def)  | 
|
1070  | 
next  | 
|
1071  | 
fix n :: nat  | 
|
1072  | 
    from sets have "{X n..b} \<in> sets M" by (cases "X n \<le> b") auto
 | 
|
1073  | 
    with int have "set_integrable M {X n..b} f"
 | 
|
1074  | 
by (rule set_integrable_subset) auto  | 
|
1075  | 
    thus "(\<lambda>x. indicat_real {X n..b} x *\<^sub>R f x) \<in> borel_measurable M"
 | 
|
1076  | 
by (simp add: set_integrable_def)  | 
|
1077  | 
qed  | 
|
1078  | 
qed  | 
|
1079  | 
||
| 
73253
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1080  | 
|
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1081  | 
|
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1082  | 
theorem integral_Markov_inequality':  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1083  | 
fixes u :: "'a \<Rightarrow> real"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1084  | 
assumes [measurable]: "set_integrable M A u" and "A \<in> sets M"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1085  | 
assumes "AE x in M. x \<in> A \<longrightarrow> u x \<ge> 0" and "0 < (c::real)"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1086  | 
  shows "emeasure M {x\<in>A. u x \<ge> c} \<le> (1/c::real) * (\<integral>x\<in>A. u x \<partial>M)"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1087  | 
proof -  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1088  | 
have "(\<lambda>x. u x * indicator A x) \<in> borel_measurable M"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1089  | 
using assms by (auto simp: set_integrable_def mult_ac)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1090  | 
hence "(\<lambda>x. ennreal (u x * indicator A x)) \<in> borel_measurable M"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1091  | 
by measurable  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1092  | 
also have "(\<lambda>x. ennreal (u x * indicator A x)) = (\<lambda>x. ennreal (u x) * indicator A x)"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1093  | 
by (intro ext) (auto simp: indicator_def)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1094  | 
finally have meas: "\<dots> \<in> borel_measurable M" .  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1095  | 
from assms(3) have AE: "AE x in M. 0 \<le> u x * indicator A x"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1096  | 
by eventually_elim (auto simp: indicator_def)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1097  | 
have nonneg: "set_lebesgue_integral M A u \<ge> 0"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1098  | 
unfolding set_lebesgue_integral_def  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1099  | 
by (intro Bochner_Integration.integral_nonneg_AE eventually_mono[OF AE]) (auto simp: mult_ac)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1100  | 
|
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1101  | 
have A: "A \<subseteq> space M"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1102  | 
using \<open>A \<in> sets M\<close> by (simp add: sets.sets_into_space)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1103  | 
  have "{x \<in> A. u x \<ge> c} = {x \<in> A. ennreal(1/c) * u x \<ge> 1}"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1104  | 
using \<open>c>0\<close> A by (auto simp: ennreal_mult'[symmetric])  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1105  | 
  then have "emeasure M {x \<in> A. u x \<ge> c} = emeasure M ({x \<in> A. ennreal(1/c) * u x \<ge> 1})"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1106  | 
by simp  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1107  | 
also have "... \<le> ennreal(1/c) * (\<integral>\<^sup>+ x. ennreal(u x) * indicator A x \<partial>M)"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1108  | 
by (intro nn_integral_Markov_inequality meas assms)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1109  | 
also have "(\<integral>\<^sup>+ x. ennreal(u x) * indicator A x \<partial>M) = ennreal (set_lebesgue_integral M A u)"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1110  | 
unfolding set_lebesgue_integral_def nn_integral_set_ennreal using assms AE  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1111  | 
by (subst nn_integral_eq_integral) (simp_all add: mult_ac set_integrable_def)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1112  | 
finally show ?thesis  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1113  | 
using \<open>c > 0\<close> nonneg by (subst ennreal_mult) auto  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1114  | 
qed  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1115  | 
|
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1116  | 
theorem integral_Markov_inequality'_measure:  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1117  | 
assumes [measurable]: "set_integrable M A u" and "A \<in> sets M"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1118  | 
and "AE x in M. x \<in> A \<longrightarrow> 0 \<le> u x" "0 < (c::real)"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1119  | 
  shows "measure M {x\<in>A. u x \<ge> c} \<le> (\<integral>x\<in>A. u x \<partial>M) / c"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1120  | 
proof -  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1121  | 
have nonneg: "set_lebesgue_integral M A u \<ge> 0"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1122  | 
unfolding set_lebesgue_integral_def  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1123  | 
by (intro Bochner_Integration.integral_nonneg_AE eventually_mono[OF assms(3)])  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1124  | 
(auto simp: mult_ac)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1125  | 
  have le: "emeasure M {x\<in>A. u x \<ge> c} \<le> ennreal ((1/c) * (\<integral>x\<in>A. u x \<partial>M))"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1126  | 
by (rule integral_Markov_inequality') (use assms in auto)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1127  | 
also have "\<dots> < top"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1128  | 
by auto  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1129  | 
  finally have "ennreal (measure M {x\<in>A. u x \<ge> c}) = emeasure M {x\<in>A. u x \<ge> c}"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1130  | 
by (intro emeasure_eq_ennreal_measure [symmetric]) auto  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1131  | 
also note le  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1132  | 
finally show ?thesis using nonneg  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1133  | 
by (subst (asm) ennreal_le_iff)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1134  | 
(auto intro!: divide_nonneg_pos Bochner_Integration.integral_nonneg_AE assms)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1135  | 
qed  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1136  | 
|
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1137  | 
theorem%important (in finite_measure) Chernoff_ineq_ge:  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1138  | 
assumes s: "s > 0"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1139  | 
assumes integrable: "set_integrable M A (\<lambda>x. exp (s * f x))" and "A \<in> sets M"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1140  | 
  shows   "measure M {x\<in>A. f x \<ge> a} \<le> exp (-s * a) * (\<integral>x\<in>A. exp (s * f x) \<partial>M)"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1141  | 
proof -  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1142  | 
  have "{x\<in>A. f x \<ge> a} = {x\<in>A. exp (s * f x) \<ge> exp (s * a)}"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1143  | 
using s by auto  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1144  | 
also have "measure M \<dots> \<le> set_lebesgue_integral M A (\<lambda>x. exp (s * f x)) / exp (s * a)"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1145  | 
by (intro integral_Markov_inequality'_measure assms) auto  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1146  | 
finally show ?thesis  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1147  | 
by (simp add: exp_minus field_simps)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1148  | 
qed  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1149  | 
|
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1150  | 
theorem%important (in finite_measure) Chernoff_ineq_le:  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1151  | 
assumes s: "s > 0"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1152  | 
assumes integrable: "set_integrable M A (\<lambda>x. exp (-s * f x))" and "A \<in> sets M"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1153  | 
  shows   "measure M {x\<in>A. f x \<le> a} \<le> exp (s * a) * (\<integral>x\<in>A. exp (-s * f x) \<partial>M)"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1154  | 
proof -  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1155  | 
  have "{x\<in>A. f x \<le> a} = {x\<in>A. exp (-s * f x) \<ge> exp (-s * a)}"
 | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1156  | 
using s by auto  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1157  | 
also have "measure M \<dots> \<le> set_lebesgue_integral M A (\<lambda>x. exp (-s * f x)) / exp (-s * a)"  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1158  | 
by (intro integral_Markov_inequality'_measure assms) auto  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1159  | 
finally show ?thesis  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1160  | 
by (simp add: exp_minus field_simps)  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1161  | 
qed  | 
| 
 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 
Manuel Eberl <eberlm@in.tum.de> 
parents: 
70721 
diff
changeset
 | 
1162  | 
|
| 
59092
 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 
hoelzl 
parents:  
diff
changeset
 | 
1163  | 
end  |