| author | desharna | 
| Mon, 10 Jun 2024 13:44:46 +0200 | |
| changeset 80321 | 31b9dfbe534c | 
| parent 80133 | e414bcc5a39e | 
| child 80768 | c7723cc15de8 | 
| 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: 
63958diff
changeset | 3 | Author: Sébastien Gouëzel sebastien.gouezel@univ-rennes1.fr | 
| 80133 
e414bcc5a39e
Acknowledgement of Ata Keskin for his Martingales material
 paulson <lp15@cam.ac.uk> parents: 
79950diff
changeset | 4 | Author: Ata Keskin, TU Muenchen | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 5 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 6 | 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 | 7 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 8 | 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 | 9 | *) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 10 | |
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 11 | chapter \<open>Integrals over a Set\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 12 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 13 | theory Set_Integral | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 14 | imports Radon_Nikodym | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 15 | begin | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 16 | |
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 17 | section \<open>Notation\<close> | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 18 | |
| 70136 | 19 | 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 | 20 | |
| 70136 | 21 | 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 | 22 | |
| 70136 | 23 | 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 | 24 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 25 | syntax | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 26 | "_ascii_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" | 
| 79630 
a8407aa7f916
Further adjustments to the syntax for Lebesgue integration
 paulson <lp15@cam.ac.uk> parents: 
79599diff
changeset | 27 |   ("(4LINT (_):(_)/|(_)./ _)" [0,0,0,10] 10)
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 28 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 29 | translations | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 30 | "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 | 31 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 32 | (* | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 33 | 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 | 34 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 35 | LBINT x. f | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 36 | LBINT x : A. f | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 37 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 38 | 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 | 39 | *) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 40 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 41 | syntax | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 42 | "_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real" | 
| 79630 
a8407aa7f916
Further adjustments to the syntax for Lebesgue integration
 paulson <lp15@cam.ac.uk> parents: 
79599diff
changeset | 43 |   ("(2LBINT _./ _)" [0,10] 10)
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 44 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 45 | syntax | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 46 | "_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> real" | 
| 79630 
a8407aa7f916
Further adjustments to the syntax for Lebesgue integration
 paulson <lp15@cam.ac.uk> parents: 
79599diff
changeset | 47 |   ("(3LBINT _:_./ _)" [0,0,10] 10)
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 48 | |
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 49 | section \<open>Basic properties\<close> | 
| 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 | (* | 
| 61945 | 52 | 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 | 53 | 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 | 54 | *) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 55 | |
| 68721 | 56 | lemma set_integrable_cong: | 
| 57 | assumes "M = M'" "A = A'" "\<And>x. x \<in> A \<Longrightarrow> f x = f' x" | |
| 58 | shows "set_integrable M A f = set_integrable M' A' f'" | |
| 59 | proof - | |
| 60 | have "(\<lambda>x. indicator A x *\<^sub>R f x) = (\<lambda>x. indicator A' x *\<^sub>R f' x)" | |
| 73536 | 61 | using assms by (auto simp: indicator_def of_bool_def) | 
| 68721 | 62 | thus ?thesis by (simp add: set_integrable_def assms) | 
| 63 | qed | |
| 64 | ||
| 62083 | 65 | lemma set_borel_measurable_sets: | 
| 66 | fixes f :: "_ \<Rightarrow> _::real_normed_vector" | |
| 67 | assumes "set_borel_measurable M X f" "B \<in> sets borel" "X \<in> sets M" | |
| 68 | shows "f -` B \<inter> X \<in> sets M" | |
| 69 | proof - | |
| 70 | 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: 
67399diff
changeset | 71 | using assms unfolding set_borel_measurable_def by (subst borel_measurable_restrict_space_iff) auto | 
| 62083 | 72 | then have "f -` B \<inter> space (restrict_space M X) \<in> sets (restrict_space M X)" | 
| 73 | by (rule measurable_sets) fact | |
| 74 | with \<open>X \<in> sets M\<close> show ?thesis | |
| 75 | by (subst (asm) sets_restrict_space_iff) (auto simp: space_restrict_space) | |
| 76 | qed | |
| 77 | ||
| 77322 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 78 | lemma set_integrable_bound: | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 79 |   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 80 |     and g :: "'a \<Rightarrow> 'c::{banach, second_countable_topology}"
 | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 81 | assumes "set_integrable M A f" "set_borel_measurable M A g" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 82 | assumes "AE x in M. x \<in> A \<longrightarrow> norm (g x) \<le> norm (f x)" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 83 | shows "set_integrable M A g" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 84 | unfolding set_integrable_def | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 85 | proof (rule Bochner_Integration.integrable_bound) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 86 | from assms(1) show "integrable M (\<lambda>x. indicator A x *\<^sub>R f x)" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 87 | by (simp add: set_integrable_def) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 88 | from assms(2) show "(\<lambda>x. indicat_real A x *\<^sub>R g x) \<in> borel_measurable M" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 89 | by (simp add: set_borel_measurable_def) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 90 | from assms(3) show "AE x in M. norm (indicat_real A x *\<^sub>R g x) \<le> norm (indicat_real A x *\<^sub>R f x)" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 91 | by eventually_elim (simp add: indicator_def) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 92 | qed | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 93 | |
| 67977 
557ea2740125
Probability builds with new definitions
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 94 | 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: 
67974diff
changeset | 95 | by (auto simp: set_lebesgue_integral_def) | 
| 
557ea2740125
Probability builds with new definitions
 paulson <lp15@cam.ac.uk> parents: 
67974diff
changeset | 96 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 97 | lemma set_lebesgue_integral_cong: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 98 | 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 | 99 | 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: 
67399diff
changeset | 100 | unfolding set_lebesgue_integral_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 101 | using assms | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 102 | 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 | 103 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 104 | 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 | 105 | 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 | 106 | assumes "AE x \<in> A in M. f x = g x" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 107 | shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)" | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 108 | proof- | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 109 | 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 | 110 | using assms by auto | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 111 | thus ?thesis | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 112 | 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 | 113 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 114 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 115 | lemma set_integrable_cong_AE: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 116 | "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: 
59867diff
changeset | 117 | 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 | 118 | 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: 
67399diff
changeset | 119 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 120 | 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 | 121 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 122 | lemma set_integrable_subset: | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 123 |   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: 
59867diff
changeset | 124 | 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 | 125 | 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 | 126 | proof - | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 127 | 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: 
67399diff
changeset | 128 | using assms integrable_mult_indicator set_integrable_def by blast | 
| 61808 | 129 | with \<open>B \<subseteq> A\<close> show ?thesis | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 130 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 131 | 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 | 132 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 133 | |
| 67339 | 134 | lemma set_integrable_restrict_space: | 
| 135 |   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | |
| 136 | assumes f: "set_integrable M S f" and T: "T \<in> sets (restrict_space M S)" | |
| 137 | shows "set_integrable M T f" | |
| 138 | proof - | |
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 139 | 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: 
67399diff
changeset | 140 | using T by (auto simp: sets_restrict_space) | 
| 67339 | 141 | 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: 
67399diff
changeset | 142 | using \<open>T' \<in> sets M\<close> f integrable_mult_indicator set_integrable_def by blast | 
| 67339 | 143 | then show ?thesis | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 144 | unfolding set_integrable_def | 
| 67339 | 145 | unfolding T_eq indicator_inter_arith by (simp add: ac_simps) | 
| 146 | qed | |
| 147 | ||
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 148 | (* 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 | 149 | (* 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 | 150 | |
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 151 | lemma set_integral_scaleR_left: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 152 | assumes "A \<in> sets M" "c \<noteq> 0 \<Longrightarrow> integrable M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 153 | shows "(LINT t:A|M. f t *\<^sub>R c) = (LINT t:A|M. f t) *\<^sub>R c" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 154 | unfolding set_lebesgue_integral_def | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 155 | using integrable_mult_indicator[OF assms] | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 156 | by (subst integral_scaleR_left[symmetric], auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 157 | |
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 158 | 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: 
67399diff
changeset | 159 | unfolding set_lebesgue_integral_def | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63627diff
changeset | 160 | 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 | 161 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 162 | 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 | 163 |   fixes a :: "'a::{real_normed_field, second_countable_topology}"
 | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 164 | 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: 
67399diff
changeset | 165 | unfolding set_lebesgue_integral_def | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63627diff
changeset | 166 | 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 | 167 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 168 | 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 | 169 |   fixes a :: "'a::{real_normed_field, second_countable_topology}"
 | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 170 | 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: 
67399diff
changeset | 171 | unfolding set_lebesgue_integral_def | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63627diff
changeset | 172 | 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 | 173 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 174 | lemma set_integral_divide_zero [simp]: | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59358diff
changeset | 175 |   fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
 | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 176 | 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: 
67399diff
changeset | 177 | unfolding set_lebesgue_integral_def | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63627diff
changeset | 178 | 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 | 179 | (auto split: split_indicator) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 180 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 181 | 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 | 182 | 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: 
67399diff
changeset | 183 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 184 | 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 | 185 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 186 | 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 | 187 |   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 | 188 | 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: 
67399diff
changeset | 189 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 190 | 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 | 191 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 192 | 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 | 193 |   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 | 194 | 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: 
67399diff
changeset | 195 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 196 | 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 | 197 | |
| 70721 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 198 | 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: 
70365diff
changeset | 199 |   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: 
70365diff
changeset | 200 | assumes "a \<noteq> 0" | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 201 | 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: 
70365diff
changeset | 202 | proof | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 203 | 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: 
70365diff
changeset | 204 | 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: 
70365diff
changeset | 205 | 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: 
70365diff
changeset | 206 | 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: 
70365diff
changeset | 207 | using assms by auto | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 208 | qed auto | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 209 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 210 | 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 | 211 |   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 | 212 | 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: 
67399diff
changeset | 213 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 214 | 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 | 215 | |
| 70721 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 216 | 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: 
70365diff
changeset | 217 |   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: 
70365diff
changeset | 218 | assumes "a \<noteq> 0" | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 219 | 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: 
70365diff
changeset | 220 | 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: 
70365diff
changeset | 221 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 222 | lemma set_integrable_divide [simp, intro]: | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59358diff
changeset | 223 |   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 | 224 | 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 | 225 | 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 | 226 | proof - | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 227 | 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: 
67399diff
changeset | 228 | 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 | 229 | 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 | 230 | by (auto split: split_indicator) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 231 | finally show ?thesis | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 232 | unfolding set_integrable_def . | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 233 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 234 | |
| 70721 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 235 | 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: 
70365diff
changeset | 236 |   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: 
70365diff
changeset | 237 | assumes "a \<noteq> 0" | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 238 | 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: 
70365diff
changeset | 239 | by (simp add: divide_inverse assms) | 
| 
47258727fa42
A few new theorems, tidying up and deletion of obsolete material
 paulson <lp15@cam.ac.uk> parents: 
70365diff
changeset | 240 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 241 | 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 | 242 |   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 | 243 | 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 | 244 | shows "set_integrable M A (\<lambda>x. f x + g x)" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 245 | 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: 
67399diff
changeset | 246 | 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 | 247 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 248 | 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 | 249 | assumes "set_integrable M A f" "set_integrable M A g" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 250 | shows "set_integrable M A (\<lambda>x. f x - g x)" 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: 
67399diff
changeset | 251 | 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 | 252 | |
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 253 | 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: 
67399diff
changeset | 254 | 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 | 255 | 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 | 256 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 257 | lemma set_integral_complex_of_real: | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 258 | "(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: 
67399diff
changeset | 259 | 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 | 260 | by (subst integral_complex_of_real[symmetric]) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63627diff
changeset | 261 | (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 | 262 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 263 | lemma set_integral_mono: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 264 | fixes f g :: "_ \<Rightarrow> real" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 265 | 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 | 266 | "\<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 | 267 | 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: 
67399diff
changeset | 268 | 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: 
67399diff
changeset | 269 | 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 | 270 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 271 | 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 | 272 | fixes f g :: "_ \<Rightarrow> real" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 273 | 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 | 274 | "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 | 275 | 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: 
67399diff
changeset | 276 | 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: 
67399diff
changeset | 277 | 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 | 278 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 279 | 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: 
67399diff
changeset | 280 | 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 | 281 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 282 | lemma set_integrable_abs_iff: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 283 | fixes f :: "_ \<Rightarrow> real" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 284 | 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: 
67399diff
changeset | 285 | 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 | 286 | 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 | 287 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 288 | lemma set_integrable_abs_iff': | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 289 | fixes f :: "_ \<Rightarrow> real" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
59867diff
changeset | 290 | 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 | 291 | 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: 
67399diff
changeset | 292 | 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 | 293 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 294 | lemma set_integrable_discrete_difference: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 295 |   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 | 296 | assumes "countable X" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 297 | 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 | 298 |   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 | 299 | 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: 
67399diff
changeset | 300 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 301 | 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 | 302 | 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 | 303 | 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 | 304 | qed fact+ | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 305 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 306 | lemma set_integral_discrete_difference: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 307 |   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 | 308 | assumes "countable X" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 309 | 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 | 310 |   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 | 311 | 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: 
67399diff
changeset | 312 | 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 | 313 | 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 | 314 | 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 | 315 | 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 | 316 | qed fact+ | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 317 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 318 | lemma set_integrable_Un: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 319 |   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 | 320 | 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 | 321 | 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 | 322 | 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 | 323 | proof - | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 324 | 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 | 325 | 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: 
67399diff
changeset | 326 | 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: 
67399diff
changeset | 327 | unfolding set_integrable_def using integrable_add by blast | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 328 | then show ?thesis | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 329 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 330 | 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 | 331 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 332 | |
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 333 | lemma set_integrable_empty [simp]: "set_integrable M {} f"
 | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 334 | by (auto simp: set_integrable_def) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 335 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 336 | lemma set_integrable_UN: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 337 |   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 | 338 | 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 | 339 | "\<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 | 340 | 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: 
67399diff
changeset | 341 | using assms | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 342 | 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 | 343 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 344 | lemma set_integral_Un: | 
| 
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 "A \<inter> B = {}"
 | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 347 | 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 | 348 | and "set_integrable M B f" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 349 | shows "(LINT x:A\<union>B|M. f x) = (LINT x:A|M. f x) + (LINT x:B|M. f x)" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 350 | using assms | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 351 | unfolding set_integrable_def set_lebesgue_integral_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 352 | 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 | 353 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 354 | lemma set_integral_cong_set: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 355 |   fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 356 | 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 | 357 | and ae: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 358 | 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: 
67399diff
changeset | 359 | 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 | 360 | proof (rule integral_cong_AE) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 361 | 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 | 362 | 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: 
67399diff
changeset | 363 | 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 | 364 | |
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 365 | 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 | 366 |   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 | 367 | 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 | 368 | shows "set_borel_measurable M B f" | 
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 369 | proof- | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 370 | 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: 
67399diff
changeset | 371 | using assms unfolding set_borel_measurable_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 372 | using borel_measurable_indicator borel_measurable_scaleR by blast | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 373 | 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 | 374 | 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: 
67399diff
changeset | 375 | ultimately show ?thesis | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 376 | 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 | 377 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 378 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 379 | lemma set_integral_Un_AE: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 380 |   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 | 381 | 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 | 382 | 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 | 383 | and "set_integrable M B f" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 384 | shows "(LINT x:A\<union>B|M. f x) = (LINT x:A|M. f x) + (LINT x:B|M. f x)" | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 385 | proof - | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 386 | 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 | 387 | 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 | 388 | 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: 
67399diff
changeset | 389 | using integrable_iff_bounded set_borel_measurable_def set_integrable_def by blast | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 390 | 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: 
59867diff
changeset | 391 | 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 | 392 | 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 | 393 | using ae by auto | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 394 | 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 | 395 | 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 | 396 | qed fact | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 397 | 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 | 398 | 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 | 399 | 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 | 400 | using ae | 
| 67399 | 401 | 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 | 402 | (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 | 403 | finally show ?thesis . | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 404 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 405 | |
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 406 | 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 | 407 |   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 | 408 | 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 | 409 | 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 | 410 | 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 | 411 | using assms | 
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 412 | proof induction | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 413 | case (insert x F) | 
| 69313 | 414 |   then have "A x \<inter> \<Union>(A ` F) = {}"
 | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 415 | by (meson disjoint_family_on_insert) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 416 | with insert show ?case | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 417 | 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: 
67399diff
changeset | 418 | 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 | 419 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 420 | (* TODO: find a better name? *) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 421 | lemma pos_integrable_to_top: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 422 | fixes l::real | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 423 | 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 | 424 | 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 | 425 | and intgbl: "\<And>i::nat. set_integrable M (A i) f" | 
| 61969 | 426 | 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: 
67399diff
changeset | 427 | shows "set_integrable M (\<Union>i. A i) f" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 428 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 429 | apply (rule integrable_monotone_convergence[where f = "\<lambda>i::nat. \<lambda>x. indicator (A i) x *\<^sub>R f x" and x = l]) | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 430 | apply (rule intgbl [unfolded set_integrable_def]) | 
| 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 431 | prefer 3 apply (rule lim [unfolded set_lebesgue_integral_def]) | 
| 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 432 | apply (rule AE_I2) | 
| 61808 | 433 | 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 | 434 | proof (rule AE_I2) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 435 |   { fix x assume "x \<in> space M"
 | 
| 61969 | 436 | 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 | 437 | proof cases | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 438 | assume "\<exists>i. x \<in> A i" | 
| 74362 | 439 | then obtain i where "x \<in> A i" .. | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 440 | then have "\<forall>\<^sub>F i in sequentially. x \<in> A i" | 
| 61808 | 441 | using \<open>x \<in> A i\<close> \<open>mono A\<close> by (auto simp: eventually_sequentially mono_def) | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 442 | with eventually_mono have "\<forall>\<^sub>F i in sequentially. indicat_real (A i) x *\<^sub>R f x = indicat_real (\<Union> (range A)) x *\<^sub>R f x" | 
| 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 443 | by fastforce | 
| 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 444 | then show ?thesis | 
| 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 445 | by (intro tendsto_eventually) | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 446 | qed auto } | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 447 | 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: 
62083diff
changeset | 448 | apply (rule borel_measurable_LIMSEQ_real) | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 449 | apply assumption | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 450 | 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 | 451 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 452 | |
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 453 | text \<open>Proof from Royden, \emph{Real Analysis}, p. 91.\<close>
 | 
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 454 | 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 | 455 |   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 | 456 | 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 | 457 |     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 | 458 | and intgbl: "set_integrable M (\<Union>i. A i) f" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 459 | 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: 
67399diff
changeset | 460 | unfolding set_lebesgue_integral_def | 
| 70136 | 461 | proof (subst integral_suminf[symmetric]) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 462 | 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: 
67399diff
changeset | 463 | using intgbl unfolding set_integrable_def [symmetric] | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 464 | 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 | 465 |   { 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 | 466 | 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 | 467 | 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 | 468 | note sums = this | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 469 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 470 | 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: 
67399diff
changeset | 471 | 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 | 472 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 473 | 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 | 474 | 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 | 475 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 476 | 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 | 477 | proof (rule summableI_nonneg_bounded) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 478 | fix n | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 479 | 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 | 480 | 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: 
59867diff
changeset | 481 | |
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 482 | 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: 
67399diff
changeset | 483 | 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 | 484 | 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 | 485 | using norm_f | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 486 | 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 | 487 | 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: 
67399diff
changeset | 488 | 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: 
67399diff
changeset | 489 | unfolding set_lebesgue_integral_def set_integrable_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 490 |       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: 
67399diff
changeset | 491 | apply (auto split: split_indicator) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 492 | done | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 493 | 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 | 494 | 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 | 495 | by simp | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 496 | qed | 
| 69313 | 497 | 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)" | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 498 | by (metis (no_types, lifting) integral_cong sums sums_unique) | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 499 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 500 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 501 | lemma set_integral_cont_up: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 502 |   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 | 503 | 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 | 504 | and intgbl: "set_integrable M (\<Union>i. A i) f" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 505 | shows "(\<lambda>i. LINT x:(A i)|M. f x) \<longlonglongrightarrow> (LINT x:(\<Union>i. A i)|M. f x)" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 506 | 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 | 507 | 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 | 508 | 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 | 509 | using intgbl by (rule set_integrable_subset) auto | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 510 | 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: 
67399diff
changeset | 511 | using int_A integrable_iff_bounded set_integrable_def by blast | 
| 69313 | 512 | 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: 
67399diff
changeset | 513 | 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: 
67399diff
changeset | 514 | 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: 
67399diff
changeset | 515 | 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: 
67399diff
changeset | 516 | by fastforce | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 517 |   { fix x i assume "x \<in> A i"
 | 
| 61969 | 518 | 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 | 519 | by (intro filterlim_cong refl) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 520 | (fastforce simp: eventually_sequentially incseq_def subset_eq intro!: exI[of _ i]) } | 
| 61969 | 521 | 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 | 522 | 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 | 523 | 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: 
59867diff
changeset | 524 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 525 | (* 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 | 526 | lemma set_integral_cont_down: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 527 |   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 | 528 | 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 | 529 | and int0: "set_integrable M (A 0) f" | 
| 79599 
2c18ac57e92e
the syntax of Lebesgue integrals (LINT, LBINT, ∫, etc.) now requires parentheses
 paulson <lp15@cam.ac.uk> parents: 
77322diff
changeset | 530 | 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: 
67399diff
changeset | 531 | 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 | 532 | proof (rule integral_dominated_convergence) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 533 | 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 | 534 | 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: 
67399diff
changeset | 535 | 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: 
67399diff
changeset | 536 | 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: 
67399diff
changeset | 537 | 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: 
67399diff
changeset | 538 | by force | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 539 | 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 | 540 | using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def) | 
| 69313 | 541 | 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: 
67399diff
changeset | 542 | "\<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: 
67399diff
changeset | 543 | 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 | 544 | 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 | 545 | 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 | 546 |   { fix x i assume "x \<in> space M" "x \<notin> A i"
 | 
| 61969 | 547 | 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 | 548 | by (intro filterlim_cong refl) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 549 | (auto split: split_indicator simp: eventually_sequentially decseq_def intro!: exI[of _ i]) } | 
| 61969 | 550 | 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 | 551 | 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 | 552 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 553 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 554 | lemma set_integral_at_point: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 555 | fixes a :: real | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 556 |   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 | 557 |   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 | 558 |   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 | 559 | proof- | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 560 |   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 | 561 | by (intro set_lebesgue_integral_cong) simp_all | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 562 | then show ?thesis using assms | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 563 | 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 | 564 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 565 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 566 | |
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 567 | section \<open>Complex integrals\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 568 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 569 | 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 | 570 | "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 | 571 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 572 | 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 | 573 | "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 | 574 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 575 | syntax | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 576 | "_complex_lebesgue_integral" :: "pttrn \<Rightarrow> complex \<Rightarrow> 'a measure \<Rightarrow> complex" | 
| 79630 
a8407aa7f916
Further adjustments to the syntax for Lebesgue integration
 paulson <lp15@cam.ac.uk> parents: 
79599diff
changeset | 577 |  ("\<integral>\<^sup>C _. _ \<partial>_" [0,0] 110)
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 578 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 579 | translations | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 580 | "\<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 | 581 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 582 | syntax | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 583 | "_ascii_complex_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" | 
| 79630 
a8407aa7f916
Further adjustments to the syntax for Lebesgue integration
 paulson <lp15@cam.ac.uk> parents: 
79599diff
changeset | 584 |   ("(3CLINT _|_. _)" [0,0,10] 10)
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 585 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 586 | translations | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 587 | "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 | 588 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 589 | lemma complex_integrable_cnj [simp]: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 590 | "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 | 591 | proof | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 592 | 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 | 593 | 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 | 594 | by (rule integrable_cnj) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 595 | 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 | 596 | by simp | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 597 | qed simp | 
| 
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 | 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 | 600 | "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 | 601 | proof | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 602 | 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 | 603 | 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 | 604 | by (rule integrable_Re) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 605 | then show "integrable M f" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 606 | by simp | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 607 | qed simp | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 608 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 609 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 610 | 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 | 611 | "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 | 612 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 613 | 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 | 614 | "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 | 615 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 616 | syntax | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 617 | "_ascii_complex_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 618 |   ("(4CLINT _:_|_. _)" [0,0,0,10] 10)
 | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 619 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 620 | translations | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 621 | "CLINT x:A|M. f" == "CONST complex_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 | 622 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 623 | 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 | 624 |   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 | 625 |   shows "set_borel_measurable borel {a..b} f"
 | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 626 | unfolding set_borel_measurable_def | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 627 | 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 | 628 | |
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 629 | section \<open>NN Set Integrals\<close> | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 630 | |
| 64911 | 631 | text\<open>This notation is from Sébastien Gouëzel: His use is not directly in line with the | 
| 632 | 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: 
63958diff
changeset | 633 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 634 | 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: 
63958diff
changeset | 635 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 636 | syntax | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 637 | "_set_nn_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" | 
| 79630 
a8407aa7f916
Further adjustments to the syntax for Lebesgue integration
 paulson <lp15@cam.ac.uk> parents: 
79599diff
changeset | 638 | ("(\<integral>\<^sup>+((_)\<in>(_)./ _)/\<partial>_)" [0,0,0,110] 10)
 | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 639 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 640 | "_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal" | 
| 79630 
a8407aa7f916
Further adjustments to the syntax for Lebesgue integration
 paulson <lp15@cam.ac.uk> parents: 
79599diff
changeset | 641 | ("(\<integral>((_)\<in>(_)./ _)/\<partial>_)" [0,0,0,110] 10)
 | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 642 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 643 | translations | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 644 | "\<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: 
63958diff
changeset | 645 | "\<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: 
63958diff
changeset | 646 | |
| 77322 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 647 | lemma set_nn_integral_cong: | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 648 | assumes "M = M'" "A = B" "\<And>x. x \<in> space M \<inter> A \<Longrightarrow> f x = g x" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 649 | shows "set_nn_integral M A f = set_nn_integral M' B g" | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 650 | by (metis (mono_tags, lifting) IntI assms indicator_simps(2) mult_eq_0_iff nn_integral_cong) | 
| 
9c295f84d55f
Replacing  z powr of_int i  by  z powi i   and adding new material from the AFP
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 651 | |
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 652 | lemma nn_integral_disjoint_pair: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 653 | assumes [measurable]: "f \<in> borel_measurable M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 654 | "B \<in> sets M" "C \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 655 |           "B \<inter> C = {}"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 656 | 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: 
63958diff
changeset | 657 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 658 | 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: 
63958diff
changeset | 659 | 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: 
63958diff
changeset | 660 | 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: 
63958diff
changeset | 661 | by (auto split: split_indicator) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 662 | 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: 
63958diff
changeset | 663 | by simp | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 664 | also have "\<dots> = (\<integral>\<^sup>+x. f x * indicator B x \<partial>M) + (\<integral>\<^sup>+x. f x * indicator C x \<partial>M)" | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 665 | by (rule nn_integral_add) (auto simp add: assms mes pos) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 666 | finally show ?thesis by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 667 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 668 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 669 | lemma nn_integral_disjoint_pair_countspace: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 670 |   assumes "B \<inter> C = {}"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 671 | 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: 
63958diff
changeset | 672 | by (rule nn_integral_disjoint_pair) (simp_all add: assms) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 673 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 674 | lemma nn_integral_null_delta: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 675 | assumes "A \<in> sets M" "B \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 676 | "(A - B) \<union> (B - A) \<in> null_sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 677 | shows "(\<integral>\<^sup>+x \<in> A. f x \<partial>M) = (\<integral>\<^sup>+x \<in> B. f x \<partial>M)" | 
| 73536 | 678 | proof (rule nn_integral_cong_AE) | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 679 | have *: "AE a in M. a \<notin> (A - B) \<union> (B - A)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 680 | using assms(3) AE_not_in by blast | 
| 73536 | 681 | 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: 
63958diff
changeset | 682 | by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 683 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 684 | |
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 685 | proposition nn_integral_disjoint_family: | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 686 | 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: 
63958diff
changeset | 687 | and "disjoint_family B" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 688 | 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: 
63958diff
changeset | 689 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 690 | 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: 
63958diff
changeset | 691 | by (rule nn_integral_suminf) simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 692 | 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: 
63958diff
changeset | 693 | proof (cases) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 694 | assume "x \<in> (\<Union>n. B n)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 695 | then obtain n where "x \<in> B n" by blast | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 696 |     have a: "finite {n}" by simp
 | 
| 64911 | 697 | 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: 
63958diff
changeset | 698 | by (metis IntI UNIV_I empty_iff) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 699 |     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: 
63958diff
changeset | 700 |     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: 
63958diff
changeset | 701 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 702 | define h where "h = (\<lambda>i. f x * indicator (B i) x)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 703 |     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: 
63958diff
changeset | 704 |     then have "(\<Sum>i. h i) = (\<Sum>i\<in>{n}. h i)"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 705 | by (metis sums_unique[OF sums_finite[OF a]]) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 706 | then have "(\<Sum>i. h i) = h n" by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 707 | then have "(\<Sum>n. f x * indicator (B n) x) = f x * indicator (B n) x" using h_def by simp | 
| 64911 | 708 | then have "(\<Sum>n. f x * indicator (B n) x) = f x" using \<open>x \<in> B n\<close> indicator_def by simp | 
| 709 | 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: 
63958diff
changeset | 710 | next | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 711 | assume "x \<notin> (\<Union>n. B n)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 712 | 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: 
63958diff
changeset | 713 | have "(\<Sum>n. f x * indicator (B n) x) = 0" | 
| 64911 | 714 | by (simp add: \<open>\<And>n. f x * indicator (B n) x = 0\<close>) | 
| 715 | 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: 
63958diff
changeset | 716 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 717 | ultimately show ?thesis by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 718 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 719 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 720 | lemma nn_set_integral_add: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 721 | assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 722 | "A \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 723 | 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: 
63958diff
changeset | 724 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 725 | 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: 
63958diff
changeset | 726 | by (auto simp add: indicator_def intro!: nn_integral_cong) | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 727 | also have "\<dots> = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) + (\<integral>\<^sup>+x. g x * indicator A x \<partial>M)" | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 728 | apply (rule nn_integral_add) using assms(1) assms(2) by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 729 | finally show ?thesis by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 730 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 731 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 732 | lemma nn_set_integral_cong: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 733 | assumes "AE x in M. f x = g x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 734 | 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: 
63958diff
changeset | 735 | apply (rule nn_integral_cong_AE) using assms(1) by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 736 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 737 | lemma nn_set_integral_set_mono: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 738 | "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: 
63958diff
changeset | 739 | by (auto intro!: nn_integral_mono split: split_indicator) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 740 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 741 | lemma nn_set_integral_mono: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 742 | assumes [measurable]: "f \<in> borel_measurable M" "g \<in> borel_measurable M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 743 | "A \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 744 | and "AE x\<in>A in M. f x \<le> g x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 745 | 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: 
63958diff
changeset | 746 | by (auto intro!: nn_integral_mono_AE split: split_indicator simp: assms) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 747 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 748 | lemma nn_set_integral_space [simp]: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 749 | 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: 
63958diff
changeset | 750 | 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: 
63958diff
changeset | 751 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 752 | lemma nn_integral_count_compose_inj: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 753 | assumes "inj_on g A" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 754 | 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: 
63958diff
changeset | 755 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 756 | 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: 
63958diff
changeset | 757 | by (auto simp add: nn_integral_count_space_indicator[symmetric]) | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 758 | also have "\<dots> = (\<integral>\<^sup>+y. f y \<partial>count_space (g`A))" | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 759 | by (simp add: assms nn_integral_bij_count_space inj_on_imp_bij_betw) | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 760 | also have "\<dots> = (\<integral>\<^sup>+y \<in> g`A. f y \<partial>count_space UNIV)" | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 761 | by (auto simp add: nn_integral_count_space_indicator[symmetric]) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 762 | finally show ?thesis by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 763 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 764 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 765 | lemma nn_integral_count_compose_bij: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 766 | assumes "bij_betw g A B" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 767 | 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: 
63958diff
changeset | 768 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 769 | have "inj_on g A" using assms bij_betw_def by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 770 | 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: 
63958diff
changeset | 771 | by (rule nn_integral_count_compose_inj) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 772 | then show ?thesis using assms by (simp add: bij_betw_def) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 773 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 774 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 775 | lemma set_integral_null_delta: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 776 |   fixes f::"_ \<Rightarrow> _ :: {banach, second_countable_topology}"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 777 | 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: 
67399diff
changeset | 778 | and null: "(A - B) \<union> (B - A) \<in> null_sets M" | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 779 | 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: 
67399diff
changeset | 780 | proof (rule set_integral_cong_set) | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 781 | 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: 
67399diff
changeset | 782 | using null AE_not_in by blast | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 783 | then show "AE x in M. (x \<in> B) = (x \<in> A)" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 784 | by auto | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 785 | qed (simp_all add: set_borel_measurable_def) | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 786 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 787 | lemma set_integral_space: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 788 | assumes "integrable M f" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 789 | 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: 
67399diff
changeset | 790 | 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: 
63958diff
changeset | 791 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 792 | lemma null_if_pos_func_has_zero_nn_int: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 793 | fixes f::"'a \<Rightarrow> ennreal" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 794 | assumes [measurable]: "f \<in> borel_measurable M" "A \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 795 | 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: 
63958diff
changeset | 796 | shows "A \<in> null_sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 797 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 798 | have "AE x in M. f x * indicator A x = 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 799 | 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: 
63958diff
changeset | 800 | 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: 
63958diff
changeset | 801 | 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: 
63958diff
changeset | 802 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 803 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 804 | lemma null_if_pos_func_has_zero_int: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 805 | assumes [measurable]: "integrable M f" "A \<in> sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 806 | 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: 
63958diff
changeset | 807 | shows "A \<in> null_sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 808 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 809 | have "AE x in M. indicator A x * f x = 0" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 810 | 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: 
67399diff
changeset | 811 | 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: 
67399diff
changeset | 812 | by (auto simp: set_lebesgue_integral_def) | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 813 | 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: 
63958diff
changeset | 814 | 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: 
63958diff
changeset | 815 | 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: 
63958diff
changeset | 816 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 817 | |
| 64911 | 818 | text\<open>The next lemma is a variant of \<open>density_unique\<close>. Note that it uses the notation | 
| 819 | for nonnegative set integrals introduced earlier.\<close> | |
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 820 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 821 | lemma (in sigma_finite_measure) density_unique2: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 822 | assumes [measurable]: "f \<in> borel_measurable M" "f' \<in> borel_measurable M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 823 | 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: 
63958diff
changeset | 824 | shows "AE x in M. f x = f' x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 825 | proof (rule density_unique) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 826 | show "density M f = density M f'" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 827 | by (intro measure_eqI) (auto simp: emeasure_density intro!: density_eq) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 828 | qed (auto simp add: assms) | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 829 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 830 | |
| 64911 | 831 | 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: 
63958diff
changeset | 832 | using Hahn-Banach theorem and linear forms. Since they are not yet easily available, I | 
| 64911 | 833 | only formulate it for real-valued functions.\<close> | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 834 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 835 | lemma density_unique_real: | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 836 | fixes f f'::"_ \<Rightarrow> real" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 837 | assumes M[measurable]: "integrable M f" "integrable M f'" | 
| 64283 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 838 | 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: 
63958diff
changeset | 839 | shows "AE x in M. f x = f' x" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 840 | proof - | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 841 |   define A where "A = {x \<in> space M. f x < f' x}"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 842 | then have [measurable]: "A \<in> sets M" by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 843 | 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: 
67399diff
changeset | 844 | 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: 
63958diff
changeset | 845 | 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: 
63958diff
changeset | 846 | then have "A \<in> null_sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 847 | 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: 
63958diff
changeset | 848 | 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: 
63958diff
changeset | 849 | 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: 
63958diff
changeset | 850 | |
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 851 |   define B where "B = {x \<in> space M. f' x < f x}"
 | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 852 | then have [measurable]: "B \<in> sets M" by simp | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 853 | 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: 
67399diff
changeset | 854 | 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: 
63958diff
changeset | 855 | 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: 
63958diff
changeset | 856 | then have "B \<in> null_sets M" | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 857 | 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: 
63958diff
changeset | 858 | 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: 
63958diff
changeset | 859 | 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: 
63958diff
changeset | 860 | then show ?thesis using * by auto | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 861 | qed | 
| 
979cdfdf7a79
HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
 hoelzl parents: 
63958diff
changeset | 862 | |
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 863 | section \<open>Scheffé's lemma\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 864 | |
| 69566 | 865 | 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: 
64283diff
changeset | 866 | 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: 
64283diff
changeset | 867 | 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: 
64283diff
changeset | 868 | variations) are known as Scheffe lemma. | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 869 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 870 | The formalization is more painful as one should jump back and forth between reals and ereals and justify | 
| 64911 | 871 | 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: 
64283diff
changeset | 872 | |
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 873 | proposition Scheffe_lemma1: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 874 | 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: 
64283diff
changeset | 875 | "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: 
64283diff
changeset | 876 | "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: 
64283diff
changeset | 877 | 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: 
69566diff
changeset | 878 | proof - | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 879 | 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: 
64283diff
changeset | 880 | using assms(1) assms(2) by simp_all | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 881 | 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: 
64283diff
changeset | 882 | 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: 
64283diff
changeset | 883 | 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: 
64283diff
changeset | 884 | 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: 
64283diff
changeset | 885 | 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: 
64283diff
changeset | 886 | 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: 
64283diff
changeset | 887 | by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 888 | 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: 
64283diff
changeset | 889 | by (auto simp: ennreal_mult_eq_top_iff) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 890 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 891 |   {
 | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 892 | 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: 
64283diff
changeset | 893 | 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: 
64283diff
changeset | 894 | 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: 
64283diff
changeset | 895 | 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: 
64283diff
changeset | 896 | 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: 
64283diff
changeset | 897 | by (rule tendsto_add) (auto simp add: a) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 898 | 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: 
64283diff
changeset | 899 | 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: 
64283diff
changeset | 900 | 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: 
64283diff
changeset | 901 | 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: 
64283diff
changeset | 902 | 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: 
64283diff
changeset | 903 | } | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 904 | 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: 
64283diff
changeset | 905 | 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: 
64283diff
changeset | 906 | by (simp add: nn_integral_cong_AE ennreal_mult) | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 907 | also have "\<dots> = 2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M)" by (rule nn_integral_cmult) auto | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 908 | 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: 
64283diff
changeset | 909 | by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 910 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 911 | 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: 
64283diff
changeset | 912 | by (rule nn_integral_add) (auto simp add: assms) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 913 | 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: 
64283diff
changeset | 914 | 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: 
64283diff
changeset | 915 | by simp | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 916 | also have "\<dots> = (\<integral>\<^sup>+x. norm(f x) \<partial>M) + limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x) \<partial>M))" | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 917 | by (rule Limsup_const_add, auto simp add: finint) | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 918 | also have "\<dots> \<le> (\<integral>\<^sup>+x. norm(f x) \<partial>M) + (\<integral>\<^sup>+x. norm(f x) \<partial>M)" | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 919 | using assms(4) by (simp add: add_left_mono) | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 920 | also have "\<dots> = 2 * (\<integral>\<^sup>+x. norm(f x) \<partial>M)" | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 921 | unfolding one_add_one[symmetric] distrib_right by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 922 | 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: 
64283diff
changeset | 923 | 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: 
64283diff
changeset | 924 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 925 | have le: "ennreal (norm (F n x - f x)) \<le> ennreal (norm (f x)) + ennreal (norm (F n x))" for n x | 
| 68403 | 926 | 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: 
64283diff
changeset | 927 | 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: 
64283diff
changeset | 928 | by (rule nn_integral_mono) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 929 | |
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 930 | 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: 
64283diff
changeset | 931 | by (simp add: int_liminf) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 932 | 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: 
64283diff
changeset | 933 | by (rule nn_integral_liminf) auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 934 | 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: 
64283diff
changeset | 935 | 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: 
64283diff
changeset | 936 | proof (intro arg_cong[where f=liminf] ext) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 937 | fix n | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 938 | have "\<And>x. ennreal(G n x) = ennreal(norm(f x)) + ennreal(norm(F n x)) - ennreal(norm(F n x - f x))" | 
| 68403 | 939 | 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: 
64283diff
changeset | 940 | 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: 
64283diff
changeset | 941 | = (\<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: 
64283diff
changeset | 942 | proof (rule nn_integral_diff) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 943 | 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: 
64283diff
changeset | 944 | by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 945 | 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: 
64283diff
changeset | 946 | 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: 
64283diff
changeset | 947 | 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: 
64283diff
changeset | 948 | qed (auto simp add: assms) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 949 | 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: 
64283diff
changeset | 950 | by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 951 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 952 | 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: 
64283diff
changeset | 953 | 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: 
64283diff
changeset | 954 | 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: 
64283diff
changeset | 955 | by (intro add_mono) auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 956 | 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: 
64283diff
changeset | 957 | 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: 
64283diff
changeset | 958 | by (intro add_mono liminf_minus_ennreal le2) auto | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 959 | 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: 
64283diff
changeset | 960 | 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: 
64283diff
changeset | 961 | 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: 
64283diff
changeset | 962 | by fact | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 963 | 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: 
64283diff
changeset | 964 | using fin2 by simp | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 965 | then show ?thesis | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 966 | by (rule tendsto_0_if_Limsup_eq_0_ennreal) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 967 | qed | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 968 | |
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 969 | proposition Scheffe_lemma2: | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 970 |   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: 
64283diff
changeset | 971 | 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: 
64283diff
changeset | 972 | "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: 
64283diff
changeset | 973 | "\<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: 
64283diff
changeset | 974 | 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: 
69566diff
changeset | 975 | proof (rule Scheffe_lemma1) | 
| 64284 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 976 | fix n::nat | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 977 | 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: 
64283diff
changeset | 978 | 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: 
64283diff
changeset | 979 | 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: 
64283diff
changeset | 980 | qed (auto simp add: assms Limsup_bounded) | 
| 
f3b905b2eee2
HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
 hoelzl parents: 
64283diff
changeset | 981 | |
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 982 | lemma tendsto_set_lebesgue_integral_at_right: | 
| 68721 | 983 |   fixes a b :: real and f :: "real \<Rightarrow> 'a :: {banach,second_countable_topology}"
 | 
| 984 |   assumes "a < b" and sets: "\<And>a'. a' \<in> {a<..b} \<Longrightarrow> {a'..b} \<in> sets M"
 | |
| 985 |       and "set_integrable M {a<..b} f"
 | |
| 986 |   shows   "((\<lambda>a'. set_lebesgue_integral M {a'..b} f) \<longlongrightarrow>
 | |
| 987 |              set_lebesgue_integral M {a<..b} f) (at_right a)"
 | |
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 988 | proof (rule tendsto_at_right_sequentially[OF assms(1)], goal_cases) | 
| 68721 | 989 | case (1 S) | 
| 990 |   have eq: "(\<Union>n. {S n..b}) = {a<..b}"
 | |
| 991 | proof safe | |
| 992 |     fix x n assume "x \<in> {S n..b}"
 | |
| 993 |     with 1(1,2)[of n] show "x \<in> {a<..b}" by auto
 | |
| 994 | next | |
| 995 |     fix x assume "x \<in> {a<..b}"
 | |
| 996 |     with order_tendstoD[OF \<open>S \<longlonglongrightarrow> a\<close>, of x] show "x \<in> (\<Union>n. {S n..b})"
 | |
| 997 | by (force simp: eventually_at_top_linorder dest: less_imp_le) | |
| 998 | qed | |
| 999 |   have "(\<lambda>n. set_lebesgue_integral M {S n..b} f) \<longlonglongrightarrow> set_lebesgue_integral M (\<Union>n. {S n..b}) f"
 | |
| 1000 | by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le) | |
| 1001 | with eq show ?case by simp | |
| 1002 | qed | |
| 1003 | ||
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1004 | section \<open>Convergence of integrals over an interval\<close> | 
| 68721 | 1005 | |
| 1006 | text \<open> | |
| 1007 | The next lemmas relate convergence of integrals over an interval to | |
| 1008 | improper integrals. | |
| 1009 | \<close> | |
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1010 | lemma tendsto_set_lebesgue_integral_at_left: | 
| 68721 | 1011 |   fixes a b :: real and f :: "real \<Rightarrow> 'a :: {banach,second_countable_topology}"
 | 
| 1012 |   assumes "a < b" and sets: "\<And>b'. b' \<in> {a..<b} \<Longrightarrow> {a..b'} \<in> sets M"
 | |
| 1013 |       and "set_integrable M {a..<b} f"
 | |
| 1014 |   shows   "((\<lambda>b'. set_lebesgue_integral M {a..b'} f) \<longlongrightarrow>
 | |
| 1015 |              set_lebesgue_integral M {a..<b} f) (at_left b)"
 | |
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1016 | proof (rule tendsto_at_left_sequentially[OF assms(1)], goal_cases) | 
| 68721 | 1017 | case (1 S) | 
| 1018 |   have eq: "(\<Union>n. {a..S n}) = {a..<b}"
 | |
| 1019 | proof safe | |
| 1020 |     fix x n assume "x \<in> {a..S n}"
 | |
| 1021 |     with 1(1,2)[of n] show "x \<in> {a..<b}" by auto
 | |
| 1022 | next | |
| 1023 |     fix x assume "x \<in> {a..<b}"
 | |
| 1024 |     with order_tendstoD[OF \<open>S \<longlonglongrightarrow> b\<close>, of x] show "x \<in> (\<Union>n. {a..S n})"
 | |
| 1025 | by (force simp: eventually_at_top_linorder dest: less_imp_le) | |
| 1026 | qed | |
| 1027 |   have "(\<lambda>n. set_lebesgue_integral M {a..S n} f) \<longlonglongrightarrow> set_lebesgue_integral M (\<Union>n. {a..S n}) f"
 | |
| 1028 | by (rule set_integral_cont_up) (insert assms 1, auto simp: eq incseq_def decseq_def less_imp_le) | |
| 1029 | with eq show ?case by simp | |
| 1030 | qed | |
| 1031 | ||
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1032 | proposition tendsto_set_lebesgue_integral_at_top: | 
| 68721 | 1033 |   fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
 | 
| 1034 |   assumes sets: "\<And>b. b \<ge> a \<Longrightarrow> {a..b} \<in> sets M"
 | |
| 1035 |       and int: "set_integrable M {a..} f"
 | |
| 1036 |   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: 
69566diff
changeset | 1037 | proof (rule tendsto_at_topI_sequentially) | 
| 68721 | 1038 | fix X :: "nat \<Rightarrow> real" assume "filterlim X at_top sequentially" | 
| 1039 |   show "(\<lambda>n. set_lebesgue_integral M {a..X n} f) \<longlonglongrightarrow> set_lebesgue_integral M {a..} f"
 | |
| 1040 | unfolding set_lebesgue_integral_def | |
| 1041 | proof (rule integral_dominated_convergence) | |
| 1042 |     show "integrable M (\<lambda>x. indicat_real {a..} x *\<^sub>R norm (f x))"
 | |
| 1043 | using integrable_norm[OF int[unfolded set_integrable_def]] by simp | |
| 1044 |     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"
 | |
| 1045 | proof | |
| 1046 | fix x | |
| 1047 | from \<open>filterlim X at_top sequentially\<close> | |
| 1048 | have "eventually (\<lambda>n. x \<le> X n) sequentially" | |
| 1049 | unfolding filterlim_at_top_ge[where c=x] by auto | |
| 1050 |       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: 
70136diff
changeset | 1051 | by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono) | 
| 68721 | 1052 | qed | 
| 1053 |     fix n show "AE x in M. norm (indicator {a..X n} x *\<^sub>R f x) \<le> 
 | |
| 1054 |                              indicator {a..} x *\<^sub>R norm (f x)"
 | |
| 1055 | by (auto split: split_indicator) | |
| 1056 | next | |
| 1057 |     from int show "(\<lambda>x. indicat_real {a..} x *\<^sub>R f x) \<in> borel_measurable M"
 | |
| 1058 | by (simp add: set_integrable_def) | |
| 1059 | next | |
| 1060 | fix n :: nat | |
| 1061 |     from sets have "{a..X n} \<in> sets M" by (cases "X n \<ge> a") auto
 | |
| 1062 |     with int have "set_integrable M {a..X n} f"
 | |
| 1063 | by (rule set_integrable_subset) auto | |
| 1064 |     thus "(\<lambda>x. indicat_real {a..X n} x *\<^sub>R f x) \<in> borel_measurable M"
 | |
| 1065 | by (simp add: set_integrable_def) | |
| 1066 | qed | |
| 1067 | qed | |
| 1068 | ||
| 69737 
ec3cc98c38db
tagged 4 theories
 Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk> parents: 
69566diff
changeset | 1069 | proposition tendsto_set_lebesgue_integral_at_bot: | 
| 68721 | 1070 |   fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
 | 
| 1071 |   assumes sets: "\<And>a. a \<le> b \<Longrightarrow> {a..b} \<in> sets M"
 | |
| 1072 |       and int: "set_integrable M {..b} f"
 | |
| 1073 |     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: 
69566diff
changeset | 1074 | proof (rule tendsto_at_botI_sequentially) | 
| 68721 | 1075 | fix X :: "nat \<Rightarrow> real" assume "filterlim X at_bot sequentially" | 
| 1076 |   show "(\<lambda>n. set_lebesgue_integral M {X n..b} f) \<longlonglongrightarrow> set_lebesgue_integral M {..b} f"
 | |
| 1077 | unfolding set_lebesgue_integral_def | |
| 1078 | proof (rule integral_dominated_convergence) | |
| 1079 |     show "integrable M (\<lambda>x. indicat_real {..b} x *\<^sub>R norm (f x))"
 | |
| 1080 | using integrable_norm[OF int[unfolded set_integrable_def]] by simp | |
| 1081 |     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"
 | |
| 1082 | proof | |
| 1083 | fix x | |
| 1084 | from \<open>filterlim X at_bot sequentially\<close> | |
| 1085 | have "eventually (\<lambda>n. x \<ge> X n) sequentially" | |
| 1086 | unfolding filterlim_at_bot_le[where c=x] by auto | |
| 1087 |       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: 
70136diff
changeset | 1088 | by (intro tendsto_eventually) (auto split: split_indicator elim!: eventually_mono) | 
| 68721 | 1089 | qed | 
| 1090 |     fix n show "AE x in M. norm (indicator {X n..b} x *\<^sub>R f x) \<le> 
 | |
| 1091 |                              indicator {..b} x *\<^sub>R norm (f x)"
 | |
| 1092 | by (auto split: split_indicator) | |
| 1093 | next | |
| 1094 |     from int show "(\<lambda>x. indicat_real {..b} x *\<^sub>R f x) \<in> borel_measurable M"
 | |
| 1095 | by (simp add: set_integrable_def) | |
| 1096 | next | |
| 1097 | fix n :: nat | |
| 1098 |     from sets have "{X n..b} \<in> sets M" by (cases "X n \<le> b") auto
 | |
| 1099 |     with int have "set_integrable M {X n..b} f"
 | |
| 1100 | by (rule set_integrable_subset) auto | |
| 1101 |     thus "(\<lambda>x. indicat_real {X n..b} x *\<^sub>R f x) \<in> borel_measurable M"
 | |
| 1102 | by (simp add: set_integrable_def) | |
| 1103 | qed | |
| 1104 | qed | |
| 1105 | ||
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1106 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1107 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1108 | theorem integral_Markov_inequality': | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1109 | fixes u :: "'a \<Rightarrow> real" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1110 | 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: 
70721diff
changeset | 1111 | 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: 
70721diff
changeset | 1112 |   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: 
70721diff
changeset | 1113 | proof - | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1114 | 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: 
70721diff
changeset | 1115 | 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: 
70721diff
changeset | 1116 | 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: 
70721diff
changeset | 1117 | by measurable | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1118 | 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: 
70721diff
changeset | 1119 | 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: 
70721diff
changeset | 1120 | 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: 
70721diff
changeset | 1121 | 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: 
70721diff
changeset | 1122 | 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: 
70721diff
changeset | 1123 | 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: 
70721diff
changeset | 1124 | unfolding set_lebesgue_integral_def | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1125 | 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: 
70721diff
changeset | 1126 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1127 | have A: "A \<subseteq> space M" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1128 | 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: 
70721diff
changeset | 1129 |   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: 
70721diff
changeset | 1130 | 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: 
70721diff
changeset | 1131 |   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: 
70721diff
changeset | 1132 | by simp | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1133 | also have "\<dots> \<le> ennreal(1/c) * (\<integral>\<^sup>+ x. ennreal(u x) * indicator A x \<partial>M)" | 
| 73253 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1134 | 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: 
70721diff
changeset | 1135 | 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: 
70721diff
changeset | 1136 | 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: 
70721diff
changeset | 1137 | 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: 
70721diff
changeset | 1138 | finally show ?thesis | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1139 | 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: 
70721diff
changeset | 1140 | qed | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1141 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1142 | theorem integral_Markov_inequality'_measure: | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1143 | 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: 
70721diff
changeset | 1144 | 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: 
70721diff
changeset | 1145 |   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: 
70721diff
changeset | 1146 | proof - | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1147 | 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: 
70721diff
changeset | 1148 | unfolding set_lebesgue_integral_def | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1149 | 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: 
70721diff
changeset | 1150 | (auto simp: mult_ac) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1151 |   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: 
70721diff
changeset | 1152 | 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: 
70721diff
changeset | 1153 | also have "\<dots> < top" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1154 | by auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1155 |   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: 
70721diff
changeset | 1156 | 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: 
70721diff
changeset | 1157 | also note le | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1158 | finally show ?thesis using nonneg | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1159 | by (subst (asm) ennreal_le_iff) | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1160 | (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: 
70721diff
changeset | 1161 | qed | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1162 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1163 | 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: 
70721diff
changeset | 1164 | assumes s: "s > 0" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1165 | 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: 
70721diff
changeset | 1166 |   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: 
70721diff
changeset | 1167 | proof - | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1168 |   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: 
70721diff
changeset | 1169 | using s by auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1170 | 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: 
70721diff
changeset | 1171 | 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: 
70721diff
changeset | 1172 | finally show ?thesis | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1173 | 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: 
70721diff
changeset | 1174 | qed | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1175 | |
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1176 | 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: 
70721diff
changeset | 1177 | assumes s: "s > 0" | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1178 | 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: 
70721diff
changeset | 1179 |   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: 
70721diff
changeset | 1180 | proof - | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1181 |   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: 
70721diff
changeset | 1182 | using s by auto | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1183 | 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: 
70721diff
changeset | 1184 | 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: 
70721diff
changeset | 1185 | finally show ?thesis | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1186 | 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: 
70721diff
changeset | 1187 | qed | 
| 
f6bb31879698
HOL-Analysis/Probability: Hoeffding's inequality, negative binomial distribution, etc.
 Manuel Eberl <eberlm@in.tum.de> parents: 
70721diff
changeset | 1188 | |
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1189 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1190 | section \<open>Integrable Simple Functions\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1191 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1192 | text \<open>This section is from the Martingales AFP entry, by Ata Keskin, TU München\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1193 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1194 | text \<open>We restate some basic results concerning Bochner-integrable functions.\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1195 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1196 | lemma integrable_implies_simple_function_sequence: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1197 |   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1198 | assumes "integrable M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1199 | obtains s where "\<And>i. simple_function M (s i)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1200 |       and "\<And>i. emeasure M {y \<in> space M. s i y \<noteq> 0} \<noteq> \<infinity>"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1201 | and "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1202 | and "\<And>x i. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1203 | proof- | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1204 | have f: "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. norm (f x) \<partial>M) < \<infinity>" using assms unfolding integrable_iff_bounded by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1205 | obtain s where s: "\<And>i. simple_function M (s i)" "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" "\<And>i x. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)" using borel_measurable_implies_sequence_metric[OF f(1)] unfolding norm_conv_dist by metis | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1206 |   {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1207 | fix i | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1208 | have "(\<integral>\<^sup>+x. norm (s i x) \<partial>M) \<le> (\<integral>\<^sup>+x. ennreal (2 * norm (f x)) \<partial>M)" using s by (intro nn_integral_mono, auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1209 | also have "\<dots> < \<infinity>" using f by (simp add: nn_integral_cmult ennreal_mult_less_top ennreal_mult) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1210 | finally have sbi: "Bochner_Integration.simple_bochner_integrable M (s i)" using s by (intro simple_bochner_integrableI_bounded) auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1211 |     hence "emeasure M {y \<in> space M. s i y \<noteq> 0} \<noteq> \<infinity>" by (auto intro: integrableI_simple_bochner_integrable simple_bochner_integrable.cases)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1212 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1213 | thus ?thesis using that s by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1214 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1215 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1216 | text \<open>Simple functions can be represented by sums of indicator functions.\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1217 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1218 | lemma simple_function_indicator_representation_banach: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1219 |   fixes f ::"'a \<Rightarrow> 'b :: {second_countable_topology, banach}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1220 | assumes "simple_function M f" "x \<in> space M" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1221 |   shows "f x = (\<Sum>y \<in> f ` space M. indicator (f -` {y} \<inter> space M) x *\<^sub>R y)"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1222 | (is "?l = ?r") | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1223 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1224 |   have "?r = (\<Sum>y \<in> f ` space M. (if y = f x then indicator (f -` {y} \<inter> space M) x *\<^sub>R y else 0))" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1225 | by (auto intro!: sum.cong) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1226 |   also have "\<dots> =  indicator (f -` {f x} \<inter> space M) x *\<^sub>R f x" using assms by (auto dest: simple_functionD)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1227 | also have "\<dots> = f x" using assms by (auto simp: indicator_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1228 | finally show ?thesis by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1229 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1230 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1231 | lemma simple_function_indicator_representation_AE: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1232 |   fixes f ::"'a \<Rightarrow> 'b :: {second_countable_topology, banach}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1233 | assumes "simple_function M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1234 |   shows "AE x in M. f x = (\<Sum>y \<in> f ` space M. indicator (f -` {y} \<inter> space M) x *\<^sub>R y)"  
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1235 | by (metis (mono_tags, lifting) AE_I2 simple_function_indicator_representation_banach assms) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1236 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1237 | lemmas simple_function_scaleR[intro] = simple_function_compose2[where h="(*\<^sub>R)"] | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1238 | lemmas integrable_simple_function = simple_bochner_integrable.intros[THEN has_bochner_integral_simple_bochner_integrable, THEN integrable.intros] | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1239 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1240 | text \<open>Induction rule for simple integrable functions.\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1241 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1242 | lemma\<^marker>\<open>tag important\<close> integrable_simple_function_induct[consumes 2, case_names cong indicator add, induct set: simple_function]: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1243 |   fixes f :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1244 |   assumes f: "simple_function M f" "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1245 |   assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity> 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1246 |                     \<Longrightarrow> simple_function M g \<Longrightarrow> emeasure M {y \<in> space M. g y \<noteq> 0} \<noteq> \<infinity> 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1247 | \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1248 | assumes indicator: "\<And>A y. A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> P (\<lambda>x. indicator A x *\<^sub>R y)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1249 |   assumes add: "\<And>f g. simple_function M f \<Longrightarrow> emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow> 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1250 |                       simple_function M g \<Longrightarrow> emeasure M {y \<in> space M. g y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow> 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1251 | (\<And>z. z \<in> space M \<Longrightarrow> norm (f z + g z) = norm (f z) + norm (g z)) \<Longrightarrow> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1252 | P f \<Longrightarrow> P g \<Longrightarrow> P (\<lambda>x. f x + g x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1253 | shows "P f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1254 | proof- | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1255 |   let ?f = "\<lambda>x. (\<Sum>y\<in>f ` space M. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1256 | have f_ae_eq: "f x = ?f x" if "x \<in> space M" for x using simple_function_indicator_representation_banach[OF f(1) that] . | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1257 |   moreover have "emeasure M {y \<in> space M. ?f y \<noteq> 0} \<noteq> \<infinity>" by (metis (no_types, lifting) Collect_cong calculation f(2))
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1258 |   moreover have "P (\<lambda>x. \<Sum>y\<in>S. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1259 |                 "simple_function M (\<lambda>x. \<Sum>y\<in>S. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1260 |                 "emeasure M {y \<in> space M. (\<Sum>x\<in>S. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} \<noteq> \<infinity>"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1261 | if "S \<subseteq> f ` space M" for S using simple_functionD(1)[OF assms(1), THEN rev_finite_subset, OF that] that | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1262 | proof (induction rule: finite_induct) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1263 | case empty | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1264 |     {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1265 | case 1 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1266 |       then show ?case using indicator[of "{}"] by force
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1267 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1268 | case 2 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1269 | then show ?case by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1270 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1271 | case 3 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1272 | then show ?case by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1273 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1274 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1275 | case (insert x F) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1276 |     have "(f -` {x} \<inter> space M) \<subseteq> {y \<in> space M. f y \<noteq> 0}" if "x \<noteq> 0" using that by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1277 |     moreover have "{y \<in> space M. f y \<noteq> 0} = space M - (f -` {0} \<inter> space M)" by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1278 |     moreover have "space M - (f -` {0} \<inter> space M) \<in> sets M" using simple_functionD(2)[OF f(1)] by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1279 |     ultimately have fin_0: "emeasure M (f -` {x} \<inter> space M) < \<infinity>" if "x \<noteq> 0" using that by (metis emeasure_mono f(2) infinity_ennreal_def top.not_eq_extremum top_unique)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1280 |     hence fin_1: "emeasure M {y \<in> space M. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x \<noteq> 0} \<noteq> \<infinity>" if "x \<noteq> 0" by (metis (mono_tags, lifting) emeasure_mono f(1) indicator_simps(2) linorder_not_less mem_Collect_eq scaleR_eq_0_iff simple_functionD(2) subsetI that)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1281 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1282 |     have *: "(\<Sum>y\<in>insert x F. indicat_real (f -` {y} \<inter> space M) xa *\<^sub>R y) = (\<Sum>y\<in>F. indicat_real (f -` {y} \<inter> space M) xa *\<^sub>R y) + indicat_real (f -` {x} \<inter> space M) xa *\<^sub>R x" for xa by (metis (no_types, lifting) Diff_empty Diff_insert0 add.commute insert.hyps(1) insert.hyps(2) sum.insert_remove)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1283 |     have **: "{y \<in> space M. (\<Sum>x\<in>insert x F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} \<subseteq> {y \<in> space M. (\<Sum>x\<in>F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} \<union> {y \<in> space M. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x \<noteq> 0}" unfolding * by fastforce    
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1284 |     {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1285 | case 1 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1286 | hence x: "x \<in> f ` space M" and F: "F \<subseteq> f ` space M" by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1287 | show ?case | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1288 | proof (cases "x = 0") | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1289 | case True | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1290 | then show ?thesis unfolding * using insert(3)[OF F] by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1291 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1292 | case False | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1293 |         have norm_argument: "norm ((\<Sum>y\<in>F. indicat_real (f -` {y} \<inter> space M) z *\<^sub>R y) + indicat_real (f -` {x} \<inter> space M) z *\<^sub>R x) = norm (\<Sum>y\<in>F. indicat_real (f -` {y} \<inter> space M) z *\<^sub>R y) + norm (indicat_real (f -` {x} \<inter> space M) z *\<^sub>R x)" if z: "z \<in> space M" for z
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1294 | proof (cases "f z = x") | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1295 | case True | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1296 |           have "indicat_real (f -` {y} \<inter> space M) z *\<^sub>R y = 0" if "y \<in> F" for y
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1297 | using True insert(2) z that 1 unfolding indicator_def by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1298 |           hence "(\<Sum>y\<in>F. indicat_real (f -` {y} \<inter> space M) z *\<^sub>R y) = 0" by (meson sum.neutral)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1299 | then show ?thesis by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1300 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1301 | case False | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1302 | then show ?thesis by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1303 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1304 | show ?thesis | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1305 | using False simple_functionD(2)[OF f(1)] insert(3,5)[OF F] simple_function_scaleR fin_0 fin_1 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1306 | by (subst *, subst add, subst simple_function_sum) (blast intro: norm_argument indicator)+ | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1307 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1308 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1309 | case 2 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1310 | hence x: "x \<in> f ` space M" and F: "F \<subseteq> f ` space M" by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1311 | show ?case | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1312 | proof (cases "x = 0") | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1313 | case True | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1314 | then show ?thesis unfolding * using insert(4)[OF F] by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1315 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1316 | case False | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1317 | then show ?thesis unfolding * using insert(4)[OF F] simple_functionD(2)[OF f(1)] by fast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1318 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1319 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1320 | case 3 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1321 | hence x: "x \<in> f ` space M" and F: "F \<subseteq> f ` space M" by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1322 | show ?case | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1323 | proof (cases "x = 0") | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1324 | case True | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1325 | then show ?thesis unfolding * using insert(5)[OF F] by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1326 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1327 | case False | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1328 |         have "emeasure M {y \<in> space M. (\<Sum>x\<in>insert x F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} \<le> emeasure M ({y \<in> space M. (\<Sum>x\<in>F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} \<union> {y \<in> space M. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x \<noteq> 0})"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1329 | using ** simple_functionD(2)[OF insert(4)[OF F]] simple_functionD(2)[OF f(1)] by (intro emeasure_mono, force+) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1330 |         also have "\<dots> \<le> emeasure M {y \<in> space M. (\<Sum>x\<in>F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} + emeasure M {y \<in> space M. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x \<noteq> 0}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1331 | using simple_functionD(2)[OF insert(4)[OF F]] simple_functionD(2)[OF f(1)] by (intro emeasure_subadditive, force+) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1332 | also have "\<dots> < \<infinity>" using insert(5)[OF F] fin_1[OF False] by (simp add: less_top) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1333 | finally show ?thesis by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1334 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1335 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1336 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1337 |   moreover have "simple_function M (\<lambda>x. \<Sum>y\<in>f ` space M. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)" using calculation by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1338 |   moreover have "P (\<lambda>x. \<Sum>y\<in>f ` space M. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)" using calculation by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1339 | ultimately show ?thesis by (intro cong[OF _ _ f(1,2)], blast, presburger+) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1340 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1341 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1342 | text \<open>Induction rule for non-negative simple integrable functions\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1343 | lemma\<^marker>\<open>tag important\<close> integrable_simple_function_induct_nn[consumes 3, case_names cong indicator add, induct set: simple_function]: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1344 |   fixes f :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1345 |   assumes f: "simple_function M f" "emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity>" "\<And>x. x \<in> space M \<longrightarrow> f x \<ge> 0"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1346 |   assumes cong: "\<And>f g. simple_function M f \<Longrightarrow> emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x \<ge> 0) \<Longrightarrow> simple_function M g \<Longrightarrow> emeasure M {y \<in> space M. g y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> g x \<ge> 0) \<Longrightarrow> (\<And>x. x \<in> space M \<Longrightarrow> f x = g x) \<Longrightarrow> P f \<Longrightarrow> P g"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1347 | assumes indicator: "\<And>A y. y \<ge> 0 \<Longrightarrow> A \<in> sets M \<Longrightarrow> emeasure M A < \<infinity> \<Longrightarrow> P (\<lambda>x. indicator A x *\<^sub>R y)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1348 |   assumes add: "\<And>f g. (\<And>x. x \<in> space M \<Longrightarrow> f x \<ge> 0) \<Longrightarrow> simple_function M f \<Longrightarrow> emeasure M {y \<in> space M. f y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow> 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1349 |                       (\<And>x. x \<in> space M \<Longrightarrow> g x \<ge> 0) \<Longrightarrow> simple_function M g \<Longrightarrow> emeasure M {y \<in> space M. g y \<noteq> 0} \<noteq> \<infinity> \<Longrightarrow> 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1350 | (\<And>z. z \<in> space M \<Longrightarrow> norm (f z + g z) = norm (f z) + norm (g z)) \<Longrightarrow> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1351 | P f \<Longrightarrow> P g \<Longrightarrow> P (\<lambda>x. f x + g x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1352 | shows "P f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1353 | proof- | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1354 |   let ?f = "\<lambda>x. (\<Sum>y\<in>f ` space M. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1355 | have f_ae_eq: "f x = ?f x" if "x \<in> space M" for x using simple_function_indicator_representation_banach[OF f(1) that] . | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1356 |   moreover have "emeasure M {y \<in> space M. ?f y \<noteq> 0} \<noteq> \<infinity>" by (metis (no_types, lifting) Collect_cong calculation f(2))
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1357 |   moreover have "P (\<lambda>x. \<Sum>y\<in>S. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1358 |                 "simple_function M (\<lambda>x. \<Sum>y\<in>S. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1359 |                 "emeasure M {y \<in> space M. (\<Sum>x\<in>S. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} \<noteq> \<infinity>"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1360 |                 "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> (\<Sum>y\<in>S. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1361 | if "S \<subseteq> f ` space M" for S using simple_functionD(1)[OF assms(1), THEN rev_finite_subset, OF that] that | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1362 | proof (induction rule: finite_subset_induct') | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1363 | case empty | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1364 |     {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1365 | case 1 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1366 |       then show ?case using indicator[of 0 "{}"] by force
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1367 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1368 | case 2 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1369 | then show ?case by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1370 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1371 | case 3 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1372 | then show ?case by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1373 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1374 | case 4 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1375 | then show ?case by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1376 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1377 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1378 | case (insert x F) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1379 |     have "(f -` {x} \<inter> space M) \<subseteq> {y \<in> space M. f y \<noteq> 0}" if "x \<noteq> 0" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1380 | using that by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1381 |     moreover have "{y \<in> space M. f y \<noteq> 0} = space M - (f -` {0} \<inter> space M)"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1382 | by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1383 |     moreover have "space M - (f -` {0} \<inter> space M) \<in> sets M" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1384 | using simple_functionD(2)[OF f(1)] by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1385 |     ultimately have fin_0: "emeasure M (f -` {x} \<inter> space M) < \<infinity>" if "x \<noteq> 0" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1386 | using that by (metis emeasure_mono f(2) infinity_ennreal_def top.not_eq_extremum top_unique) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1387 |     hence fin_1: "emeasure M {y \<in> space M. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x \<noteq> 0} \<noteq> \<infinity>" if "x \<noteq> 0" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1388 | by (metis (mono_tags, lifting) emeasure_mono f(1) indicator_simps(2) linorder_not_less mem_Collect_eq scaleR_eq_0_iff simple_functionD(2) subsetI that) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1389 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1390 | have nonneg_x: "x \<ge> 0" using insert f by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1391 |     have *: "(\<Sum>y\<in>insert x F. indicat_real (f -` {y} \<inter> space M) xa *\<^sub>R y) = (\<Sum>y\<in>F. indicat_real (f -` {y} \<inter> space M) xa *\<^sub>R y) + indicat_real (f -` {x} \<inter> space M) xa *\<^sub>R x" for xa by (metis (no_types, lifting) add.commute insert.hyps(1) insert.hyps(4) sum.insert)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1392 |     have **: "{y \<in> space M. (\<Sum>x\<in>insert x F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} \<subseteq> {y \<in> space M. (\<Sum>x\<in>F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} \<union> {y \<in> space M. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x \<noteq> 0}" unfolding * by fastforce    
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1393 |     {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1394 | case 1 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1395 | show ?case | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1396 | proof (cases "x = 0") | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1397 | case True | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1398 | then show ?thesis unfolding * using insert by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1399 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1400 | case False | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1401 |         have norm_argument: "norm ((\<Sum>y\<in>F. indicat_real (f -` {y} \<inter> space M) z *\<^sub>R y) + indicat_real (f -` {x} \<inter> space M) z *\<^sub>R x) 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1402 |               = norm (\<Sum>y\<in>F. indicat_real (f -` {y} \<inter> space M) z *\<^sub>R y) + norm (indicat_real (f -` {x} \<inter> space M) z *\<^sub>R x)" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1403 | if z: "z \<in> space M" for z | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1404 | proof (cases "f z = x") | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1405 | case True | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1406 |           have "indicat_real (f -` {y} \<inter> space M) z *\<^sub>R y = 0" if "y \<in> F" for y 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1407 | using True insert z that 1 unfolding indicator_def by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1408 |           hence "(\<Sum>y\<in>F. indicat_real (f -` {y} \<inter> space M) z *\<^sub>R y) = 0" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1409 | by (meson sum.neutral) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1410 | thus ?thesis by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1411 | qed (force) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1412 | show ?thesis using False fin_0 fin_1 f norm_argument | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1413 | by (subst *, subst add, presburger add: insert, intro insert, intro insert, fastforce simp add: indicator_def intro!: insert(2) f(3), auto intro!: indicator insert nonneg_x) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1414 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1415 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1416 | case 2 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1417 | show ?case | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1418 | proof (cases "x = 0") | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1419 | case True | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1420 | then show ?thesis unfolding * using insert by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1421 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1422 | case False | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1423 | then show ?thesis unfolding * using insert simple_functionD(2)[OF f(1)] by fast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1424 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1425 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1426 | case 3 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1427 | show ?case | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1428 | proof (cases "x = 0") | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1429 | case True | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1430 | then show ?thesis unfolding * using insert by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1431 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1432 | case False | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1433 |         have "emeasure M {y \<in> space M. (\<Sum>x\<in>insert x F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1434 |            \<le> emeasure M ({y \<in> space M. (\<Sum>x\<in>F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} \<union> {y \<in> space M. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x \<noteq> 0})"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1435 | using ** simple_functionD(2)[OF insert(6)] simple_functionD(2)[OF f(1)] insert.IH(2) by (intro emeasure_mono, blast, simp) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1436 |         also have "\<dots> \<le> emeasure M {y \<in> space M. (\<Sum>x\<in>F. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x) \<noteq> 0} + emeasure M {y \<in> space M. indicat_real (f -` {x} \<inter> space M) y *\<^sub>R x \<noteq> 0}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1437 | using simple_functionD(2)[OF insert(6)] simple_functionD(2)[OF f(1)] by (intro emeasure_subadditive, force+) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1438 | also have "\<dots> < \<infinity>" using insert(7) fin_1[OF False] by (simp add: less_top) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1439 | finally show ?thesis by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1440 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1441 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1442 | case (4 \<xi>) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1443 | thus ?case using insert nonneg_x f(3) by (auto simp add: scaleR_nonneg_nonneg intro: sum_nonneg) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1444 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1445 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1446 |   moreover have "simple_function M (\<lambda>x. \<Sum>y\<in>f ` space M. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1447 | using calculation by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1448 |   moreover have "P (\<lambda>x. \<Sum>y\<in>f ` space M. indicat_real (f -` {y} \<inter> space M) x *\<^sub>R y)" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1449 | using calculation by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1450 | moreover have "\<And>x. x \<in> space M \<Longrightarrow> 0 \<le> f x" using f(3) by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1451 | ultimately show ?thesis | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1452 | by (smt (verit) Collect_cong f(1) local.cong) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1453 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1454 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1455 | lemma finite_nn_integral_imp_ae_finite: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1456 | fixes f :: "'a \<Rightarrow> ennreal" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1457 | assumes "f \<in> borel_measurable M" "(\<integral>\<^sup>+x. f x \<partial>M) < \<infinity>" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1458 | shows "AE x in M. f x < \<infinity>" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1459 | proof (rule ccontr, goal_cases) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1460 | case 1 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1461 |   let ?A = "space M \<inter> {x. f x = \<infinity>}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1462 | have *: "emeasure M ?A > 0" using 1 assms(1) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1463 | by (metis (mono_tags, lifting) assms(2) eventually_mono infinity_ennreal_def nn_integral_noteq_infinite top.not_eq_extremum) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1464 | have "(\<integral>\<^sup>+x. f x * indicator ?A x \<partial>M) = (\<integral>\<^sup>+x. \<infinity> * indicator ?A x \<partial>M)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1465 | by (metis (mono_tags, lifting) indicator_inter_arith indicator_simps(2) mem_Collect_eq mult_eq_0_iff) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1466 | also have "\<dots> = \<infinity> * emeasure M ?A" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1467 | using assms(1) by (intro nn_integral_cmult_indicator, simp) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1468 | also have "\<dots> = \<infinity>" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1469 | using * by fastforce | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1470 | finally have "(\<integral>\<^sup>+x. f x * indicator ?A x \<partial>M) = \<infinity>" . | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1471 | moreover have "(\<integral>\<^sup>+x. f x * indicator ?A x \<partial>M) \<le> (\<integral>\<^sup>+x. f x \<partial>M)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1472 | by (intro nn_integral_mono, simp add: indicator_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1473 | ultimately show ?case using assms(2) by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1474 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1475 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1476 | text \<open>Convergence in L1-Norm implies existence of a subsequence which convergences almost everywhere. | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1477 | This lemma is easier to use than the existing one in \<^theory>\<open>HOL-Analysis.Bochner_Integration\<close>\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1478 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1479 | lemma cauchy_L1_AE_cauchy_subseq: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1480 |   fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1481 | assumes [measurable]: "\<And>n. integrable M (s n)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1482 | and "\<And>e. e > 0 \<Longrightarrow> \<exists>N. \<forall>i\<ge>N. \<forall>j\<ge>N. LINT x|M. norm (s i x - s j x) < e" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1483 | obtains r where "strict_mono r" "AE x in M. Cauchy (\<lambda>i. s (r i) x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1484 | proof- | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1485 | have "\<exists>r. \<forall>n. (\<forall>i\<ge>r n. \<forall>j\<ge> r n. LINT x|M. norm (s i x - s j x) < (1 / 2) ^ n) \<and> (r (Suc n) > r n)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1486 | proof (intro dependent_nat_choice, goal_cases) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1487 | case 1 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1488 | then show ?case using assms(2) by presburger | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1489 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1490 | case (2 x n) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1491 | obtain N where *: "LINT x|M. norm (s i x - s j x) < (1 / 2) ^ Suc n" if "i \<ge> N" "j \<ge> N" for i j | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1492 | using assms(2)[of "(1 / 2) ^ Suc n"] by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1493 |     {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1494 | fix i j assume "i \<ge> max N (Suc x)" "j \<ge> max N (Suc x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1495 | hence "integral\<^sup>L M (\<lambda>x. norm (s i x - s j x)) < (1 / 2) ^ Suc n" using * by fastforce | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1496 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1497 | then show ?case by fastforce | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1498 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1499 | then obtain r where strict_mono: "strict_mono r" and "\<forall>i\<ge>r n. \<forall>j\<ge> r n. LINT x|M. norm (s i x - s j x) < (1 / 2) ^ n" for n | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1500 | using strict_mono_Suc_iff by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1501 | hence r_is: "LINT x|M. norm (s (r (Suc n)) x - s (r n) x) < (1 / 2) ^ n" for n by (simp add: strict_mono_leD) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1502 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1503 | define g where "g = (\<lambda>n x. (\<Sum>i\<le>n. ennreal (norm (s (r (Suc i)) x - s (r i) x))))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1504 | define g' where "g' = (\<lambda>x. \<Sum>i. ennreal (norm (s (r (Suc i)) x - s (r i) x)))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1505 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1506 | have integrable_g: "(\<integral>\<^sup>+ x. g n x \<partial>M) < 2" for n | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1507 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1508 | have "(\<integral>\<^sup>+ x. g n x \<partial>M) = (\<integral>\<^sup>+ x. (\<Sum>i\<le>n. ennreal (norm (s (r (Suc i)) x - s (r i) x))) \<partial>M)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1509 | using g_def by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1510 | also have "\<dots> = (\<Sum>i\<le>n. (\<integral>\<^sup>+ x. ennreal (norm (s (r (Suc i)) x - s (r i) x)) \<partial>M))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1511 | by (intro nn_integral_sum, simp) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1512 | also have "\<dots> = (\<Sum>i\<le>n. LINT x|M. norm (s (r (Suc i)) x - s (r i) x))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1513 | unfolding dist_norm using assms(1) by (subst nn_integral_eq_integral[OF integrable_norm], auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1514 | also have "\<dots> < ennreal (\<Sum>i\<le>n. (1 / 2) ^ i)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1515 | by (intro ennreal_lessI[OF sum_pos sum_strict_mono[OF finite_atMost _ r_is]], auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1516 | also have "\<dots> \<le> ennreal 2" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1517 | unfolding sum_gp0[of "1 / 2" n] by (intro ennreal_leI, auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1518 | finally show "(\<integral>\<^sup>+ x. g n x \<partial>M) < 2" by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1519 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1520 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1521 | have integrable_g': "(\<integral>\<^sup>+ x. g' x \<partial>M) \<le> 2" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1522 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1523 | have "incseq (\<lambda>n. g n x)" for x by (intro incseq_SucI, auto simp add: g_def ennreal_leI) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1524 | hence "convergent (\<lambda>n. g n x)" for x | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1525 | unfolding convergent_def using LIMSEQ_incseq_SUP by fast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1526 | hence "(\<lambda>n. g n x) \<longlonglongrightarrow> g' x" for x | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1527 | unfolding g_def g'_def by (intro summable_iff_convergent'[THEN iffD2, THEN summable_LIMSEQ'], blast) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1528 | hence "(\<integral>\<^sup>+ x. g' x \<partial>M) = (\<integral>\<^sup>+ x. liminf (\<lambda>n. g n x) \<partial>M)" by (metis lim_imp_Liminf trivial_limit_sequentially) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1529 | also have "\<dots> \<le> liminf (\<lambda>n. \<integral>\<^sup>+ x. g n x \<partial>M)" by (intro nn_integral_liminf, simp add: g_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1530 | also have "\<dots> \<le> liminf (\<lambda>n. 2)" using integrable_g by (intro Liminf_mono) (simp add: order_le_less) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1531 | also have "\<dots> = 2" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1532 | using sequentially_bot tendsto_iff_Liminf_eq_Limsup by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1533 | finally show ?thesis . | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1534 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1535 | hence "AE x in M. g' x < \<infinity>" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1536 | by (intro finite_nn_integral_imp_ae_finite) (auto simp add: order_le_less_trans g'_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1537 | moreover have "summable (\<lambda>i. norm (s (r (Suc i)) x - s (r i) x))" if "g' x \<noteq> \<infinity>" for x | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1538 | using that unfolding g'_def by (intro summable_suminf_not_top) fastforce+ | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1539 | ultimately have ae_summable: "AE x in M. summable (\<lambda>i. s (r (Suc i)) x - s (r i) x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1540 | using summable_norm_cancel unfolding dist_norm by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1541 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1542 |   {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1543 | fix x assume "summable (\<lambda>i. s (r (Suc i)) x - s (r i) x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1544 | hence "(\<lambda>n. \<Sum>i<n. s (r (Suc i)) x - s (r i) x) \<longlonglongrightarrow> (\<Sum>i. s (r (Suc i)) x - s (r i) x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1545 | using summable_LIMSEQ by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1546 | moreover have "(\<lambda>n. (\<Sum>i<n. s (r (Suc i)) x - s (r i) x)) = (\<lambda>n. s (r n) x - s (r 0) x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1547 | using sum_lessThan_telescope by fastforce | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1548 | ultimately have "(\<lambda>n. s (r n) x - s (r 0) x) \<longlonglongrightarrow> (\<Sum>i. s (r (Suc i)) x - s (r i) x)" by argo | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1549 | hence "(\<lambda>n. s (r n) x - s (r 0) x + s (r 0) x) \<longlonglongrightarrow> (\<Sum>i. s (r (Suc i)) x - s (r i) x) + s (r 0) x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1550 | by (intro isCont_tendsto_compose[of _ "\<lambda>z. z + s (r 0) x"], auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1551 | hence "Cauchy (\<lambda>n. s (r n) x)" by (simp add: LIMSEQ_imp_Cauchy) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1552 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1553 | hence "AE x in M. Cauchy (\<lambda>i. s (r i) x)" using ae_summable by fast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1554 | thus ?thesis by (rule that[OF strict_mono(1)]) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1555 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1556 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1557 | subsection \<open>Totally Ordered Banach Spaces\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1558 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1559 | lemma integrable_max[simp, intro]: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1560 |   fixes f :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1561 | assumes fg[measurable]: "integrable M f" "integrable M g" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1562 | shows "integrable M (\<lambda>x. max (f x) (g x))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1563 | proof (rule Bochner_Integration.integrable_bound) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1564 |   {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1565 | fix x y :: 'b | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1566 | have "norm (max x y) \<le> max (norm x) (norm y)" by linarith | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1567 | also have "\<dots> \<le> norm x + norm y" by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1568 | finally have "norm (max x y) \<le> norm (norm x + norm y)" by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1569 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1570 | thus "AE x in M. norm (max (f x) (g x)) \<le> norm (norm (f x) + norm (g x))" by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1571 | qed (auto intro: Bochner_Integration.integrable_add[OF integrable_norm[OF fg(1)] integrable_norm[OF fg(2)]]) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1572 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1573 | lemma integrable_min[simp, intro]: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1574 |   fixes f :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1575 | assumes [measurable]: "integrable M f" "integrable M g" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1576 | shows "integrable M (\<lambda>x. min (f x) (g x))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1577 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1578 | have "norm (min (f x) (g x)) \<le> norm (f x) \<or> norm (min (f x) (g x)) \<le> norm (g x)" for x by linarith | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1579 | thus ?thesis | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1580 | by (intro integrable_bound[OF integrable_max[OF integrable_norm(1,1), OF assms]], auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1581 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1582 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1583 | text \<open>Restatement of \<open>integral_nonneg_AE\<close> for functions taking values in a Banach space.\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1584 | lemma integral_nonneg_AE_banach: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1585 |   fixes f :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1586 | assumes [measurable]: "f \<in> borel_measurable M" and nonneg: "AE x in M. 0 \<le> f x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1587 | shows "0 \<le> integral\<^sup>L M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1588 | proof cases | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1589 | assume integrable: "integrable M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1590 | hence max: "(\<lambda>x. max 0 (f x)) \<in> borel_measurable M" "\<And>x. 0 \<le> max 0 (f x)" "integrable M (\<lambda>x. max 0 (f x))" by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1591 | hence "0 \<le> integral\<^sup>L M (\<lambda>x. max 0 (f x))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1592 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1593 | obtain s where *: "\<And>i. simple_function M (s i)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1594 |                     "\<And>i. emeasure M {y \<in> space M. s i y \<noteq> 0} \<noteq> \<infinity>" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1595 | "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. s i x) \<longlonglongrightarrow> f x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1596 | "\<And>x i. x \<in> space M \<Longrightarrow> norm (s i x) \<le> 2 * norm (f x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1597 | using integrable_implies_simple_function_sequence[OF integrable] by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1598 | have simple: "\<And>i. simple_function M (\<lambda>x. max 0 (s i x))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1599 | using * by fast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1600 |     have "\<And>i. {y \<in> space M. max 0 (s i y) \<noteq> 0} \<subseteq> {y \<in> space M. s i y \<noteq> 0}" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1601 | unfolding max_def by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1602 |     moreover have "\<And>i. {y \<in> space M. s i y \<noteq> 0} \<in> sets M" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1603 | using * by measurable | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1604 |     ultimately have "\<And>i. emeasure M {y \<in> space M. max 0 (s i y) \<noteq> 0} \<le> emeasure M {y \<in> space M. s i y \<noteq> 0}" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1605 | by (rule emeasure_mono) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1606 |     hence **:"\<And>i. emeasure M {y \<in> space M. max 0 (s i y) \<noteq> 0} \<noteq> \<infinity>" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1607 | using *(2) by (auto intro: order.strict_trans1 simp add: top.not_eq_extremum) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1608 | have "\<And>x. x \<in> space M \<Longrightarrow> (\<lambda>i. max 0 (s i x)) \<longlonglongrightarrow> max 0 (f x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1609 | using *(3) tendsto_max by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1610 | moreover have "\<And>x i. x \<in> space M \<Longrightarrow> norm (max 0 (s i x)) \<le> norm (2 *\<^sub>R f x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1611 | using *(4) unfolding max_def by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1612 | ultimately have tendsto: "(\<lambda>i. integral\<^sup>L M (\<lambda>x. max 0 (s i x))) \<longlonglongrightarrow> integral\<^sup>L M (\<lambda>x. max 0 (f x))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1613 | using borel_measurable_simple_function simple integrable by (intro integral_dominated_convergence[OF max(1) _ integrable_norm[OF integrable_scaleR_right], of _ 2 f], blast+) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1614 |     {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1615 |       fix h :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1616 |       assume "simple_function M h" "emeasure M {y \<in> space M. h y \<noteq> 0} \<noteq> \<infinity>" "\<And>x. x \<in> space M \<longrightarrow> h x \<ge> 0"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1617 | hence *: "integral\<^sup>L M h \<ge> 0" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1618 | proof (induct rule: integrable_simple_function_induct_nn) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1619 | case (cong f g) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1620 | then show ?case using Bochner_Integration.integral_cong by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1621 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1622 | case (indicator A y) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1623 |         hence "A \<noteq> {} \<Longrightarrow> y \<ge> 0" using sets.sets_into_space by fastforce
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1624 |         then show ?case using indicator by (cases "A = {}", auto simp add: scaleR_nonneg_nonneg)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1625 | next | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1626 | case (add f g) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1627 | then show ?case by (simp add: integrable_simple_function) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1628 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1629 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1630 | thus ?thesis using LIMSEQ_le_const[OF tendsto, of 0] ** simple by fastforce | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1631 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1632 | also have "\<dots> = integral\<^sup>L M f" using nonneg by (auto intro: integral_cong_AE) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1633 | finally show ?thesis . | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1634 | qed (simp add: not_integrable_integral_eq) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1635 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1636 | lemma integral_mono_AE_banach: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1637 |   fixes f g :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1638 | assumes "integrable M f" "integrable M g" "AE x in M. f x \<le> g x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1639 | shows "integral\<^sup>L M f \<le> integral\<^sup>L M g" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1640 | using integral_nonneg_AE_banach[of "\<lambda>x. g x - f x"] assms Bochner_Integration.integral_diff[OF assms(1,2)] by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1641 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1642 | lemma integral_mono_banach: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1643 |   fixes f g :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1644 | assumes "integrable M f" "integrable M g" "\<And>x. x \<in> space M \<Longrightarrow> f x \<le> g x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1645 | shows "integral\<^sup>L M f \<le> integral\<^sup>L M g" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1646 | using integral_mono_AE_banach assms by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1647 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1648 | subsection \<open>Auxiliary Lemmas for Set Integrals\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1649 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1650 | lemma nn_set_integral_eq_set_integral: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1651 | assumes [measurable]:"integrable M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1652 | and "AE x \<in> A in M. 0 \<le> f x" "A \<in> sets M" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1653 | shows "(\<integral>\<^sup>+x\<in>A. f x \<partial>M) = (\<integral> x \<in> A. f x \<partial>M)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1654 | proof- | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1655 | have "(\<integral>\<^sup>+x. indicator A x *\<^sub>R f x \<partial>M) = (\<integral> x \<in> A. f x \<partial>M)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1656 | unfolding set_lebesgue_integral_def | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1657 | using assms(2) by (intro nn_integral_eq_integral[of _ "\<lambda>x. indicat_real A x *\<^sub>R f x"], blast intro: assms integrable_mult_indicator, fastforce) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1658 | moreover have "(\<integral>\<^sup>+x. indicator A x *\<^sub>R f x \<partial>M) = (\<integral>\<^sup>+x\<in>A. f x \<partial>M)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1659 | by (metis ennreal_0 indicator_simps(1) indicator_simps(2) mult.commute mult_1 mult_zero_left real_scaleR_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1660 | ultimately show ?thesis by argo | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1661 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1662 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1663 | lemma set_integral_restrict_space: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1664 |   fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1665 | assumes "\<Omega> \<inter> space M \<in> sets M" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1666 | shows "set_lebesgue_integral (restrict_space M \<Omega>) A f = set_lebesgue_integral M A (\<lambda>x. indicator \<Omega> x *\<^sub>R f x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1667 | unfolding set_lebesgue_integral_def | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1668 | by (subst integral_restrict_space, auto intro!: integrable_mult_indicator assms simp: mult.commute) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1669 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1670 | lemma set_integral_const: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1671 |   fixes c :: "'b::{banach, second_countable_topology}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1672 | assumes "A \<in> sets M" "emeasure M A \<noteq> \<infinity>" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1673 | shows "set_lebesgue_integral M A (\<lambda>_. c) = measure M A *\<^sub>R c" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1674 | unfolding set_lebesgue_integral_def | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1675 | using assms by (metis has_bochner_integral_indicator has_bochner_integral_integral_eq infinity_ennreal_def less_top) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1676 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1677 | lemma set_integral_mono_banach: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1678 |   fixes f g :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1679 | assumes "set_integrable M A f" "set_integrable M A g" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1680 | "\<And>x. x \<in> A \<Longrightarrow> f x \<le> g x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1681 | shows "(LINT x:A|M. f x) \<le> (LINT x:A|M. g x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1682 | using assms unfolding set_integrable_def set_lebesgue_integral_def | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1683 | by (auto intro: integral_mono_banach split: split_indicator) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1684 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1685 | lemma set_integral_mono_AE_banach: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1686 |   fixes f g :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1687 | assumes "set_integrable M A f" "set_integrable M A g" "AE x\<in>A in M. f x \<le> g x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1688 | shows "set_lebesgue_integral M A f \<le> set_lebesgue_integral M A g" using assms unfolding set_lebesgue_integral_def by (auto simp add: set_integrable_def intro!: integral_mono_AE_banach[of M "\<lambda>x. indicator A x *\<^sub>R f x" "\<lambda>x. indicator A x *\<^sub>R g x"], simp add: indicator_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1689 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1690 | subsection \<open>Integrability and Measurability of the Diameter\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1691 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1692 | context | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1693 |   fixes s :: "nat \<Rightarrow> 'a \<Rightarrow> 'b :: {second_countable_topology, banach}" and M
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1694 | assumes bounded: "\<And>x. x \<in> space M \<Longrightarrow> bounded (range (\<lambda>i. s i x))" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1695 | begin | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1696 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1697 | lemma borel_measurable_diameter: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1698 | assumes [measurable]: "\<And>i. (s i) \<in> borel_measurable M" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1699 |   shows "(\<lambda>x. diameter {s i x |i. n \<le> i}) \<in> borel_measurable M"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1700 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1701 |   have "{s i x |i. N \<le> i} \<noteq> {}" for x N by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1702 |   hence diameter_SUP: "diameter {s i x |i. N \<le> i} = (SUP (i, j) \<in> {N..} \<times> {N..}. dist (s i x) (s j x))" for x N unfolding diameter_def by (auto intro!: arg_cong[of _ _ Sup])
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1703 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1704 |   have "case_prod dist ` ({s i x |i. n \<le> i} \<times> {s i x |i. n \<le> i}) = ((\<lambda>(i, j). dist (s i x) (s j x)) ` ({n..} \<times> {n..}))" for x by fast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1705 |   hence *: "(\<lambda>x. diameter {s i x |i. n \<le> i}) =  (\<lambda>x. Sup ((\<lambda>(i, j). dist (s i x) (s j x)) ` ({n..} \<times> {n..})))" using diameter_SUP by (simp add: case_prod_beta')
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1706 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1707 |   have "bounded ((\<lambda>(i, j). dist (s i x) (s j x)) ` ({n..} \<times> {n..}))" if "x \<in> space M" for x by (rule bounded_imp_dist_bounded[OF bounded, OF that])
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1708 |   hence bdd: "bdd_above ((\<lambda>(i, j). dist (s i x) (s j x)) ` ({n..} \<times> {n..}))" if "x \<in> space M" for x using that bounded_imp_bdd_above by presburger
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1709 |   have "fst p \<in> borel_measurable M" "snd p \<in> borel_measurable M" if "p \<in> s ` {n..} \<times> s ` {n..}" for p using that by fastforce+
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1710 |   hence "(\<lambda>x. fst p x - snd p x) \<in> borel_measurable M" if "p \<in> s ` {n..} \<times> s ` {n..}" for p using that borel_measurable_diff by simp
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1711 |   hence "(\<lambda>x. case p of (f, g) \<Rightarrow> dist (f x) (g x)) \<in> borel_measurable M" if "p \<in> s ` {n..} \<times> s ` {n..}" for p unfolding dist_norm using that by measurable
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1712 |   moreover have "countable (s ` {n..} \<times> s ` {n..})" by (intro countable_SIGMA countable_image, auto)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1713 | ultimately show ?thesis unfolding * by (auto intro!: borel_measurable_cSUP bdd) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1714 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1715 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1716 | lemma integrable_bound_diameter: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1717 | fixes f :: "'a \<Rightarrow> real" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1718 | assumes "integrable M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1719 | and [measurable]: "\<And>i. (s i) \<in> borel_measurable M" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1720 | and "\<And>x i. x \<in> space M \<Longrightarrow> norm (s i x) \<le> f x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1721 |     shows "integrable M (\<lambda>x. diameter {s i x |i. n \<le> i})"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1722 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1723 |   have "{s i x |i. N \<le> i} \<noteq> {}" for x N by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1724 |   hence diameter_SUP: "diameter {s i x |i. N \<le> i} = (SUP (i, j) \<in> {N..} \<times> {N..}. dist (s i x) (s j x))" for x N unfolding diameter_def by (auto intro!: arg_cong[of _ _ Sup])
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1725 |   {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1726 | fix x assume x: "x \<in> space M" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1727 |     let ?S = "(\<lambda>(i, j). dist (s i x) (s j x)) ` ({n..} \<times> {n..})"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1728 |     have "case_prod dist ` ({s i x |i. n \<le> i} \<times> {s i x |i. n \<le> i}) = (\<lambda>(i, j). dist (s i x) (s j x)) ` ({n..} \<times> {n..})" by fast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1729 |     hence *: "diameter {s i x |i. n \<le> i} =  Sup ?S" using diameter_SUP by (simp add: case_prod_beta')
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1730 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1731 | have "bounded ?S" by (rule bounded_imp_dist_bounded[OF bounded[OF x]]) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1732 | hence Sup_S_nonneg:"0 \<le> Sup ?S" by (auto intro!: cSup_upper2 x bounded_imp_bdd_above) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1733 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1734 | have "dist (s i x) (s j x) \<le> 2 * f x" for i j by (intro dist_triangle2[THEN order_trans, of _ 0]) (metis norm_conv_dist assms(3) x add_mono mult_2) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1735 | hence "\<forall>c \<in> ?S. c \<le> 2 * f x" by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1736 | hence "Sup ?S \<le> 2 * f x" by (intro cSup_least, auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1737 | hence "norm (Sup ?S) \<le> 2 * norm (f x)" using Sup_S_nonneg by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1738 | also have "\<dots> = norm (2 *\<^sub>R f x)" by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1739 |     finally have "norm (diameter {s i x |i. n \<le> i}) \<le> norm (2 *\<^sub>R f x)" unfolding * .
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1740 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1741 |   hence "AE x in M. norm (diameter {s i x |i. n \<le> i}) \<le> norm (2 *\<^sub>R f x)" by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1742 |   thus  "integrable M (\<lambda>x. diameter {s i x |i. n \<le> i})" using borel_measurable_diameter by (intro Bochner_Integration.integrable_bound[OF assms(1)[THEN integrable_scaleR_right[of 2]]], measurable)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1743 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1744 | end | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1745 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1746 | subsection \<open>Averaging Theorem\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1747 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1748 | text \<open>We aim to lift results from the real case to arbitrary Banach spaces. | 
| 79950 | 1749 | Our fundamental tool in this regard will be the averaging theorem. The proof of this theorem is due to Serge Lang (Real and Functional Analysis) \<^cite>\<open>"Lang_1993"\<close>. | 
| 79787 | 1750 | The theorem allows us to make statements about a function's value almost everywhere, depending on the value its integral takes on various sets of the measure space.\<close> | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1751 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1752 | text \<open>Before we introduce and prove the averaging theorem, we will first show the following lemma which is crucial for our proof. | 
| 79950 | 1753 | While not stated exactly in this manner, our proof makes use of the characterization of second-countable topological spaces given in the book General Topology by Ryszard Engelking (Theorem 4.1.15) \<^cite>\<open>"engelking_1989"\<close>. | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1754 |     (Engelking's book \emph{General Topology})\<close>
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1755 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1756 | lemma balls_countable_basis: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1757 |   obtains D :: "'a :: {metric_space, second_countable_topology} set" 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1758 |   where "topological_basis (case_prod ball ` (D \<times> (\<rat> \<inter> {0<..})))"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1759 | and "countable D" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1760 |     and "D \<noteq> {}"    
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1761 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1762 |   obtain D :: "'a set" where dense_subset: "countable D" "D \<noteq> {}" "\<lbrakk>open U; U \<noteq> {}\<rbrakk> \<Longrightarrow> \<exists>y \<in> D. y \<in> U" for U using countable_dense_exists by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1763 |   have "topological_basis (case_prod ball ` (D \<times> (\<rat> \<inter> {0<..})))"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1764 | proof (intro topological_basis_iff[THEN iffD2], fast, clarify) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1765 | fix U and x :: 'a assume asm: "open U" "x \<in> U" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1766 | obtain e where e: "e > 0" "ball x e \<subseteq> U" using asm openE by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1767 | obtain y where y: "y \<in> D" "y \<in> ball x (e / 3)" using dense_subset(3)[OF open_ball, of x "e / 3"] centre_in_ball[THEN iffD2, OF divide_pos_pos[OF e(1), of 3]] by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1768 |     obtain r where r: "r \<in> \<rat> \<inter> {e/3<..<e/2}" unfolding Rats_def using of_rat_dense[OF divide_strict_left_mono[OF _ e(1)], of 2 3] by auto
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1769 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1770 | have *: "x \<in> ball y r" using r y by (simp add: dist_commute) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1771 | hence "ball y r \<subseteq> U" using r by (intro order_trans[OF _ e(2)], simp, metric) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1772 |     moreover have "ball y r \<in> (case_prod ball ` (D \<times> (\<rat> \<inter> {0<..})))" using y(1) r by force
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1773 |     ultimately show "\<exists>B'\<in>(case_prod ball ` (D \<times> (\<rat> \<inter> {0<..}))). x \<in> B' \<and> B' \<subseteq> U" using * by meson
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1774 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1775 | thus ?thesis using that dense_subset by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1776 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1777 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1778 | context sigma_finite_measure | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1779 | begin | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1780 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1781 | text \<open>To show statements concerning \<open>\<sigma>\<close>-finite measure spaces, one usually shows the statement for finite measure spaces and uses a limiting argument to show it for the \<open>\<sigma>\<close>-finite case. | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1782 | The following induction scheme formalizes this.\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1783 | lemma sigma_finite_measure_induct[case_names finite_measure, consumes 0]: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1784 | assumes "\<And>(N :: 'a measure) \<Omega>. finite_measure N | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1785 | \<Longrightarrow> N = restrict_space M \<Omega> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1786 | \<Longrightarrow> \<Omega> \<in> sets M | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1787 | \<Longrightarrow> emeasure N \<Omega> \<noteq> \<infinity> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1788 | \<Longrightarrow> emeasure N \<Omega> \<noteq> 0 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1789 | \<Longrightarrow> almost_everywhere N Q" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1790 | and [measurable]: "Measurable.pred M Q" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1791 | shows "almost_everywhere M Q" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1792 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1793 | have *: "almost_everywhere N Q" if "finite_measure N" "N = restrict_space M \<Omega>" "\<Omega> \<in> sets M" "emeasure N \<Omega> \<noteq> \<infinity>" for N \<Omega> using that by (cases "emeasure N \<Omega> = 0", auto intro: emeasure_0_AE assms(1)) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1794 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1795 | obtain A :: "nat \<Rightarrow> 'a set" where A: "range A \<subseteq> sets M" "(\<Union>i. A i) = space M" and emeasure_finite: "emeasure M (A i) \<noteq> \<infinity>" for i using sigma_finite by metis | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1796 | note A(1)[measurable] | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1797 | have space_restr: "space (restrict_space M (A i)) = A i" for i unfolding space_restrict_space by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1798 |   {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1799 | fix i | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1800 |     have *: "{x \<in> A i \<inter> space M. Q x} = {x \<in> space M. Q x} \<inter> (A i)" by fast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1801 | have "Measurable.pred (restrict_space M (A i)) Q" using A by (intro measurableI, auto simp add: space_restr intro!: sets_restrict_space_iff[THEN iffD2], measurable, auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1802 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1803 | note this[measurable] | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1804 |   {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1805 | fix i | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1806 | have "finite_measure (restrict_space M (A i))" using emeasure_finite by (intro finite_measureI, subst space_restr, subst emeasure_restrict_space, auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1807 |     hence "emeasure (restrict_space M (A i)) {x \<in> A i. \<not>Q x} = 0" using emeasure_finite by (intro AE_iff_measurable[THEN iffD1, OF _ _ *], measurable, subst space_restr[symmetric], intro sets.top, auto simp add: emeasure_restrict_space)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1808 |     hence "emeasure M {x \<in> A i. \<not> Q x} = 0" by (subst emeasure_restrict_space[symmetric], auto)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1809 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1810 |   hence "emeasure M (\<Union>i. {x \<in> A i. \<not> Q x}) = 0" by (intro emeasure_UN_eq_0, auto)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1811 |   moreover have "(\<Union>i. {x \<in> A i. \<not> Q x}) = {x \<in> space M. \<not> Q x}" using A by auto
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1812 | ultimately show ?thesis by (intro AE_iff_measurable[THEN iffD2], auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1813 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1814 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1815 | text \<open>Real Functional Analysis - Lang\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1816 | text \<open>The Averaging Theorem allows us to make statements concerning how a function behaves almost everywhere, depending on its behaviour on average.\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1817 | lemma averaging_theorem: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1818 |   fixes f::"_ \<Rightarrow> 'b::{second_countable_topology, banach}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1819 | assumes [measurable]: "integrable M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1820 | and closed: "closed S" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1821 | and "\<And>A. A \<in> sets M \<Longrightarrow> measure M A > 0 \<Longrightarrow> (1 / measure M A) *\<^sub>R set_lebesgue_integral M A f \<in> S" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1822 | shows "AE x in M. f x \<in> S" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1823 | proof (induct rule: sigma_finite_measure_induct) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1824 | case (finite_measure N \<Omega>) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1825 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1826 | interpret finite_measure N by (rule finite_measure) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1827 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1828 | have integrable[measurable]: "integrable N f" using assms finite_measure by (auto simp: integrable_restrict_space integrable_mult_indicator) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1829 | have average: "(1 / Sigma_Algebra.measure N A) *\<^sub>R set_lebesgue_integral N A f \<in> S" if "A \<in> sets N" "measure N A > 0" for A | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1830 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1831 | have *: "A \<in> sets M" using that by (simp add: sets_restrict_space_iff finite_measure) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1832 | have "A = A \<inter> \<Omega>" by (metis finite_measure(2,3) inf.orderE sets.sets_into_space space_restrict_space that(1)) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1833 | hence "set_lebesgue_integral N A f = set_lebesgue_integral M A f" unfolding finite_measure by (subst set_integral_restrict_space, auto simp add: finite_measure set_lebesgue_integral_def indicator_inter_arith[symmetric]) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1834 | moreover have "measure N A = measure M A" using that by (auto intro!: measure_restrict_space simp add: finite_measure sets_restrict_space_iff) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1835 | ultimately show ?thesis using that * assms(3) by presburger | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1836 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1837 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1838 |   obtain D :: "'b set" where balls_basis: "topological_basis (case_prod ball ` (D \<times> (\<rat> \<inter> {0<..})))" and countable_D: "countable D" using balls_countable_basis by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1839 |   have countable_balls: "countable (case_prod ball ` (D \<times> (\<rat> \<inter> {0<..})))" using countable_rat countable_D by blast
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1840 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1841 |   obtain B where B_balls: "B \<subseteq> case_prod ball ` (D \<times> (\<rat> \<inter> {0<..}))" "\<Union>B = -S" using topological_basis[THEN iffD1, OF balls_basis] open_Compl[OF assms(2)] by meson
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1842 | hence countable_B: "countable B" using countable_balls countable_subset by fast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1843 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1844 |   define b where "b = from_nat_into (B \<union> {{}})"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1845 |   have "B \<union> {{}} \<noteq> {}" by simp
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1846 |   have range_b: "range b = B \<union> {{}}" using countable_B by (auto simp add: b_def intro!: range_from_nat_into)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1847 |   have open_b: "open (b i)" for i unfolding b_def using B_balls open_ball from_nat_into[of "B \<union> {{}}" i] by force
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1848 | have Union_range_b: "\<Union>(range b) = -S" using B_balls range_b by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1849 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1850 |   {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1851 | fix v r assume ball_in_Compl: "ball v r \<subseteq> -S" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1852 | define A where "A = f -` ball v r \<inter> space N" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1853 | have dist_less: "dist (f x) v < r" if "x \<in> A" for x using that unfolding A_def vimage_def by (simp add: dist_commute) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1854 | hence AE_less: "AE x \<in> A in N. norm (f x - v) < r" by (auto simp add: dist_norm) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1855 | have *: "A \<in> sets N" unfolding A_def by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1856 | have "emeasure N A = 0" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1857 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1858 |       {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1859 | assume asm: "emeasure N A > 0" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1860 | hence measure_pos: "measure N A > 0" unfolding emeasure_eq_measure by simp | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1861 | hence "(1 / measure N A) *\<^sub>R set_lebesgue_integral N A f - v | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1862 | = (1 / measure N A) *\<^sub>R set_lebesgue_integral N A (\<lambda>x. f x - v)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1863 | using integrable integrable_const * | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1864 | by (subst set_integral_diff(2), auto simp add: set_integrable_def set_integral_const[OF *] algebra_simps intro!: integrable_mult_indicator) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1865 | moreover have "norm (\<integral>x\<in>A. (f x - v)\<partial>N) \<le> (\<integral>x\<in>A. norm (f x - v)\<partial>N)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1866 | using * by (auto intro!: integral_norm_bound[of N "\<lambda>x. indicator A x *\<^sub>R (f x - v)", THEN order_trans] integrable_mult_indicator integrable simp add: set_lebesgue_integral_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1867 | ultimately have "norm ((1 / measure N A) *\<^sub>R set_lebesgue_integral N A f - v) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1868 | \<le> set_lebesgue_integral N A (\<lambda>x. norm (f x - v)) / measure N A" using asm by (auto intro: divide_right_mono) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1869 | also have "\<dots> < set_lebesgue_integral N A (\<lambda>x. r) / measure N A" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1870 | unfolding set_lebesgue_integral_def | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1871 | using asm * integrable integrable_const AE_less measure_pos | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1872 | by (intro divide_strict_right_mono integral_less_AE[of _ _ A] integrable_mult_indicator) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1873 | (fastforce simp add: dist_less dist_norm indicator_def)+ | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1874 | also have "\<dots> = r" using * measure_pos by (simp add: set_integral_const) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1875 | finally have "dist ((1 / measure N A) *\<^sub>R set_lebesgue_integral N A f) v < r" by (subst dist_norm) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1876 | hence "False" using average[OF * measure_pos] by (metis ComplD dist_commute in_mono mem_ball ball_in_Compl) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1877 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1878 | thus ?thesis by fastforce | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1879 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1880 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1881 | note * = this | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1882 |   {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1883 | fix b' assume "b' \<in> B" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1884 | hence ball_subset_Compl: "b' \<subseteq> -S" and ball_radius_pos: "\<exists>v \<in> D. \<exists>r>0. b' = ball v r" using B_balls by (blast, fast) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1885 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1886 | note ** = this | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1887 |   hence "emeasure N (f -` b i \<inter> space N) = 0" for i by (cases "b i = {}", simp) (metis UnE singletonD  * range_b[THEN eq_refl, THEN range_subsetD])
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1888 | hence "emeasure N (\<Union>i. f -` b i \<inter> space N) = 0" using open_b by (intro emeasure_UN_eq_0) fastforce+ | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1889 | moreover have "(\<Union>i. f -` b i \<inter> space N) = f -` (\<Union>(range b)) \<inter> space N" by blast | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1890 | ultimately have "emeasure N (f -` (-S) \<inter> space N) = 0" using Union_range_b by argo | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1891 | hence "AE x in N. f x \<notin> -S" using open_Compl[OF assms(2)] by (intro AE_iff_measurable[THEN iffD2], auto) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1892 | thus ?case by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1893 | qed (simp add: pred_sets2[OF borel_closed] assms(2)) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1894 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1895 | lemma density_zero: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1896 |   fixes f::"'a \<Rightarrow> 'b::{second_countable_topology, banach}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1897 | assumes "integrable M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1898 | and density_0: "\<And>A. A \<in> sets M \<Longrightarrow> set_lebesgue_integral M A f = 0" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1899 | shows "AE x in M. f x = 0" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1900 |   using averaging_theorem[OF assms(1), of "{0}"] assms(2)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1901 | by (simp add: scaleR_nonneg_nonneg) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1902 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1903 | text \<open>The following lemma shows that densities are unique in Banach spaces.\<close> | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1904 | lemma density_unique_banach: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1905 |   fixes f f'::"'a \<Rightarrow> 'b::{second_countable_topology, banach}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1906 | assumes "integrable M f" "integrable M f'" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1907 | and density_eq: "\<And>A. A \<in> sets M \<Longrightarrow> set_lebesgue_integral M A f = set_lebesgue_integral M A f'" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1908 | shows "AE x in M. f x = f' x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1909 | proof- | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1910 |   { 
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1911 | fix A assume asm: "A \<in> sets M" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1912 | hence "LINT x|M. indicat_real A x *\<^sub>R (f x - f' x) = 0" using density_eq assms(1,2) by (simp add: set_lebesgue_integral_def algebra_simps Bochner_Integration.integral_diff[OF integrable_mult_indicator(1,1)]) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1913 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1914 | thus ?thesis using density_zero[OF Bochner_Integration.integrable_diff[OF assms(1,2)]] by (simp add: set_lebesgue_integral_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1915 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1916 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1917 | lemma density_nonneg: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1918 |   fixes f::"_ \<Rightarrow> 'b::{second_countable_topology, banach, linorder_topology, ordered_real_vector}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1919 | assumes "integrable M f" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1920 | and "\<And>A. A \<in> sets M \<Longrightarrow> set_lebesgue_integral M A f \<ge> 0" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1921 | shows "AE x in M. f x \<ge> 0" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1922 |   using averaging_theorem[OF assms(1), of "{0..}", OF closed_atLeast] assms(2)
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1923 | by (simp add: scaleR_nonneg_nonneg) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1924 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1925 | corollary integral_nonneg_eq_0_iff_AE_banach: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1926 |   fixes f :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1927 | assumes f[measurable]: "integrable M f" and nonneg: "AE x in M. 0 \<le> f x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1928 | shows "integral\<^sup>L M f = 0 \<longleftrightarrow> (AE x in M. f x = 0)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1929 | proof | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1930 | assume *: "integral\<^sup>L M f = 0" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1931 |   {
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1932 | fix A assume asm: "A \<in> sets M" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1933 | have "0 \<le> integral\<^sup>L M (\<lambda>x. indicator A x *\<^sub>R f x)" using nonneg by (subst integral_zero[of M, symmetric], intro integral_mono_AE_banach integrable_mult_indicator asm f integrable_zero, auto simp add: indicator_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1934 | moreover have "\<dots> \<le> integral\<^sup>L M f" using nonneg by (intro integral_mono_AE_banach integrable_mult_indicator asm f, auto simp add: indicator_def) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1935 | ultimately have "set_lebesgue_integral M A f = 0" unfolding set_lebesgue_integral_def using * by force | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1936 | } | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1937 | thus "AE x in M. f x = 0" by (intro density_zero f, blast) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1938 | qed (auto simp add: integral_eq_zero_AE) | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1939 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1940 | corollary integral_eq_mono_AE_eq_AE: | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1941 |   fixes f g :: "'a \<Rightarrow> 'b :: {second_countable_topology, banach, linorder_topology, ordered_real_vector}"
 | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1942 | assumes "integrable M f" "integrable M g" "integral\<^sup>L M f = integral\<^sup>L M g" "AE x in M. f x \<le> g x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1943 | shows "AE x in M. f x = g x" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1944 | proof - | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1945 | define h where "h = (\<lambda>x. g x - f x)" | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1946 | have "AE x in M. h x = 0" unfolding h_def using assms by (subst integral_nonneg_eq_0_iff_AE_banach[symmetric]) auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1947 | then show ?thesis unfolding h_def by auto | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1948 | qed | 
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1949 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1950 | end | 
| 79772 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1951 | |
| 
817d33f8aa7f
Moving valuable library material from Martingales into the distribution
 paulson <lp15@cam.ac.uk> parents: 
79630diff
changeset | 1952 | end |