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