src/HOL/Analysis/Set_Integral.thy
author wenzelm
Tue, 01 Jan 2019 21:47:27 +0100
changeset 69566 c41954ee87cf
parent 69313 b021008c5397
child 69737 ec3cc98c38db
permissions -rw-r--r--
more antiquotations -- less LaTeX macros;
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
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     4
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     5
Notation and useful facts for working with integrals over a set.
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     6
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     7
TODO: keep all these? Need unicode translations as well.
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     8
*)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     9
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    10
theory Set_Integral
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
    11
  imports Radon_Nikodym
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    12
begin
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    13
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
    14
(*
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    15
    Notation
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    16
*)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    17
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
    18
definition%important "set_borel_measurable M A f \<equiv> (\<lambda>x. indicator A x *\<^sub>R f x) \<in> borel_measurable M"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    19
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
    20
definition%important  "set_integrable M A f \<equiv> integrable M (\<lambda>x. indicator A x *\<^sub>R f x)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    21
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
    22
definition%important  "set_lebesgue_integral M A f \<equiv> lebesgue_integral M (\<lambda>x. indicator A x *\<^sub>R f x)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    23
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    24
syntax
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    25
  "_ascii_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    26
  ("(4LINT (_):(_)/|(_)./ _)" [0,60,110,61] 60)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    27
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    28
translations
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    29
  "LINT x:A|M. f" == "CONST set_lebesgue_integral M A (\<lambda>x. f)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    30
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    31
(*
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    32
    Notation for integration wrt lebesgue measure on the reals:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    33
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    34
      LBINT x. f
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    35
      LBINT x : A. f
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    36
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    37
    TODO: keep all these? Need unicode.
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    38
*)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    39
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    40
syntax
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    41
  "_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    42
  ("(2LBINT _./ _)" [0,60] 60)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    43
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    44
syntax
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    45
  "_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> real"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    46
  ("(3LBINT _:_./ _)" [0,60,61] 60)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    47
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
    48
(*
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
    49
    Basic properties
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    50
*)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    51
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    52
(*
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61880
diff changeset
    53
lemma indicator_abs_eq: "\<And>A x. \<bar>indicator A x\<bar> = ((indicator A x) :: real)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    54
  by (auto simp add: indicator_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    55
*)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    56
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
    57
lemma set_integrable_cong:
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
    58
  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
    59
  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
    60
proof -
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
    61
  have "(\<lambda>x. indicator A x *\<^sub>R f x) = (\<lambda>x. indicator A' x *\<^sub>R f' x)"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
    62
    using assms by (auto simp: indicator_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
    63
  thus ?thesis by (simp add: set_integrable_def assms)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
    64
qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
    65
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61969
diff changeset
    66
lemma set_borel_measurable_sets:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61969
diff changeset
    67
  fixes f :: "_ \<Rightarrow> _::real_normed_vector"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61969
diff changeset
    68
  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
    69
  shows "f -` B \<inter> X \<in> sets M"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61969
diff changeset
    70
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61969
diff changeset
    71
  have "f \<in> borel_measurable (restrict_space M X)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    72
    using assms unfolding set_borel_measurable_def by (subst borel_measurable_restrict_space_iff) auto
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61969
diff changeset
    73
  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
    74
    by (rule measurable_sets) fact
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61969
diff changeset
    75
  with \<open>X \<in> sets M\<close> show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61969
diff changeset
    76
    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
    77
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61969
diff changeset
    78
67977
557ea2740125 Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents: 67974
diff changeset
    79
lemma set_lebesgue_integral_zero [simp]: "set_lebesgue_integral M A (\<lambda>x. 0) = 0"
557ea2740125 Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents: 67974
diff changeset
    80
  by (auto simp: set_lebesgue_integral_def)
557ea2740125 Probability builds with new definitions
paulson <lp15@cam.ac.uk>
parents: 67974
diff changeset
    81
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    82
lemma set_lebesgue_integral_cong:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    83
  assumes "A \<in> sets M" and "\<forall>x. x \<in> A \<longrightarrow> f x = g x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    84
  shows "(LINT x:A|M. f x) = (LINT x:A|M. g x)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    85
  unfolding set_lebesgue_integral_def
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    86
  using assms
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    87
  by (metis indicator_simps(2) real_vector.scale_zero_left)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    88
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    89
lemma set_lebesgue_integral_cong_AE:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    90
  assumes [measurable]: "A \<in> sets M" "f \<in> borel_measurable M" "g \<in> borel_measurable M"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    91
  assumes "AE x \<in> A in M. f x = g x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    92
  shows "LINT x:A|M. f x = LINT x:A|M. g x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    93
proof-
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    94
  have "AE x in M. indicator A x *\<^sub>R f x = indicator A x *\<^sub>R g x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    95
    using assms by auto
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    96
  thus ?thesis
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
    97
  unfolding set_lebesgue_integral_def by (intro integral_cong_AE) auto
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    98
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    99
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   100
lemma set_integrable_cong_AE:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   101
    "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   102
    AE x \<in> A in M. f x = g x \<Longrightarrow> A \<in> sets M \<Longrightarrow>
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   103
    set_integrable M A f = set_integrable M A g"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   104
  unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   105
  by (rule integrable_cong_AE) auto
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   106
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   107
lemma set_integrable_subset:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   108
  fixes M A B and f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   109
  assumes "set_integrable M A f" "B \<in> sets M" "B \<subseteq> A"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   110
  shows "set_integrable M B f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   111
proof -
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   112
  have "set_integrable M B (\<lambda>x. indicator A x *\<^sub>R f x)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   113
    using assms integrable_mult_indicator set_integrable_def by blast
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60615
diff changeset
   114
  with \<open>B \<subseteq> A\<close> show ?thesis
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   115
    unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   116
    by (simp add: indicator_inter_arith[symmetric] Int_absorb2)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   117
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   118
67339
d91b9d22305b HOL-Analysis: add set_integrable_restrict_space
hoelzl
parents: 66456
diff changeset
   119
lemma set_integrable_restrict_space:
d91b9d22305b HOL-Analysis: add set_integrable_restrict_space
hoelzl
parents: 66456
diff changeset
   120
  fixes f :: "'a \<Rightarrow> 'b::{banach, second_countable_topology}"
d91b9d22305b HOL-Analysis: add set_integrable_restrict_space
hoelzl
parents: 66456
diff changeset
   121
  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
   122
  shows "set_integrable M T f"
d91b9d22305b HOL-Analysis: add set_integrable_restrict_space
hoelzl
parents: 66456
diff changeset
   123
proof -
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   124
  obtain T' where T_eq: "T = S \<inter> T'" and "T' \<in> sets M" 
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   125
    using T by (auto simp: sets_restrict_space)
67339
d91b9d22305b HOL-Analysis: add set_integrable_restrict_space
hoelzl
parents: 66456
diff changeset
   126
  have \<open>integrable M (\<lambda>x. indicator T' x *\<^sub>R (indicator S x *\<^sub>R f x))\<close>
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   127
    using \<open>T' \<in> sets M\<close> f integrable_mult_indicator set_integrable_def by blast
67339
d91b9d22305b HOL-Analysis: add set_integrable_restrict_space
hoelzl
parents: 66456
diff changeset
   128
  then show ?thesis
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   129
    unfolding set_integrable_def
67339
d91b9d22305b HOL-Analysis: add set_integrable_restrict_space
hoelzl
parents: 66456
diff changeset
   130
    unfolding T_eq indicator_inter_arith by (simp add: ac_simps)
d91b9d22305b HOL-Analysis: add set_integrable_restrict_space
hoelzl
parents: 66456
diff changeset
   131
qed
d91b9d22305b HOL-Analysis: add set_integrable_restrict_space
hoelzl
parents: 66456
diff changeset
   132
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   133
(* TODO: integral_cmul_indicator should be named set_integral_const *)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   134
(* TODO: borel_integrable_atLeastAtMost should be named something like set_integrable_Icc_isCont *)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   135
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   136
lemma set_integral_scaleR_right [simp]: "LINT t:A|M. a *\<^sub>R f t = a *\<^sub>R (LINT t:A|M. f t)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   137
  unfolding set_lebesgue_integral_def
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   138
  by (subst integral_scaleR_right[symmetric]) (auto intro!: Bochner_Integration.integral_cong)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   139
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   140
lemma set_integral_mult_right [simp]:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   141
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   142
  shows "LINT t:A|M. a * f t = a * (LINT t:A|M. f t)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   143
  unfolding set_lebesgue_integral_def
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   144
  by (subst integral_mult_right_zero[symmetric]) auto
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   145
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   146
lemma set_integral_mult_left [simp]:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   147
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   148
  shows "LINT t:A|M. f t * a = (LINT t:A|M. f t) * a"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   149
  unfolding set_lebesgue_integral_def
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   150
  by (subst integral_mult_left_zero[symmetric]) auto
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   151
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   152
lemma set_integral_divide_zero [simp]:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59358
diff changeset
   153
  fixes a :: "'a::{real_normed_field, field, second_countable_topology}"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   154
  shows "LINT t:A|M. f t / a = (LINT t:A|M. f t) / a"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   155
  unfolding set_lebesgue_integral_def
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   156
  by (subst integral_divide_zero[symmetric], intro Bochner_Integration.integral_cong)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   157
     (auto split: split_indicator)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   158
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   159
lemma set_integrable_scaleR_right [simp, intro]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   160
  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. a *\<^sub>R f t)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   161
  unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   162
  unfolding scaleR_left_commute by (rule integrable_scaleR_right)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   163
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   164
lemma set_integrable_scaleR_left [simp, intro]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   165
  fixes a :: "_ :: {banach, second_countable_topology}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   166
  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. f t *\<^sub>R a)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   167
  unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   168
  using integrable_scaleR_left[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   169
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   170
lemma set_integrable_mult_right [simp, intro]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   171
  fixes a :: "'a::{real_normed_field, second_countable_topology}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   172
  shows "(a \<noteq> 0 \<Longrightarrow> set_integrable M A f) \<Longrightarrow> set_integrable M A (\<lambda>t. a * f t)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   173
  unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   174
  using integrable_mult_right[of a M "\<lambda>x. indicator A x *\<^sub>R f x"] by simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   175
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   176
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
   177
  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
   178
  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
   179
  unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   180
  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
   181
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   182
lemma set_integrable_divide [simp, intro]:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59358
diff changeset
   183
  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
   184
  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
   185
  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
   186
proof -
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   187
  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
   188
    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
   189
  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
   190
    by (auto split: split_indicator)
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   191
  finally show ?thesis 
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
qed
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_integral_add [simp, intro]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   196
  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
   197
  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
   198
  shows "set_integrable M A (\<lambda>x. f x + g x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   199
    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
   200
  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
   201
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   202
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
   203
  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
   204
  shows "set_integrable M A (\<lambda>x. f x - g x)" and "LINT x:A|M. f x - g x =
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   205
    (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
   206
  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
   207
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   208
(* question: why do we have this for negation, but multiplication by a constant
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   209
   requires an integrability assumption? *)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   210
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
   211
  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
   212
  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
   213
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   214
lemma set_integral_complex_of_real:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   215
  "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
   216
  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
   217
  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
   218
     (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
   219
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   220
lemma set_integral_mono:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   221
  fixes f g :: "_ \<Rightarrow> real"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   222
  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
   223
    "\<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
   224
  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
   225
  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
   226
  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
   227
60615
e5fa1d5d3952 Useful lemmas. The theorem concerning swapping the variables in a double integral.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   228
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
   229
  fixes f g :: "_ \<Rightarrow> real"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   230
  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
   231
    "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
   232
  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
   233
  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
   234
  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
   235
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   236
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
   237
  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
   238
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   239
lemma set_integrable_abs_iff:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   240
  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
   241
  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
   242
  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
   243
  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
   244
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   245
lemma set_integrable_abs_iff':
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   246
  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
   247
  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
   248
    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
   249
  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
   250
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   251
lemma set_integrable_discrete_difference:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   252
  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
   253
  assumes "countable X"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   254
  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
   255
  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
   256
  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
   257
  unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   258
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
   259
  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
   260
    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
   261
qed fact+
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   262
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   263
lemma set_integral_discrete_difference:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   264
  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
   265
  assumes "countable X"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   266
  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
   267
  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
   268
  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
   269
  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
   270
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
   271
  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
   272
    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
   273
qed fact+
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   274
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   275
lemma set_integrable_Un:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   276
  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
   277
  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
   278
    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
   279
  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
   280
proof -
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   281
  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
   282
    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
   283
  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
   284
    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
   285
  then show ?thesis
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   286
    unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   287
    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
   288
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   289
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   290
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
   291
  by (auto simp: set_integrable_def)
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   292
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   293
lemma set_integrable_UN:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   294
  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
   295
  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
   296
    "\<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
   297
  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
   298
  using assms
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   299
  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
   300
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   301
lemma set_integral_Un:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   302
  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
   303
  assumes "A \<inter> B = {}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   304
  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
   305
  and "set_integrable M B f"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   306
shows "LINT x:A\<union>B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   307
  using assms
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   308
  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
   309
  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
   310
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   311
lemma set_integral_cong_set:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   312
  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
   313
  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
   314
    and ae: "AE x in M. x \<in> A \<longleftrightarrow> x \<in> B"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   315
  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
   316
  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
   317
proof (rule integral_cong_AE)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   318
  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
   319
    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
   320
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
   321
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   322
lemma%important set_borel_measurable_subset:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   323
  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
   324
  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
   325
  shows "set_borel_measurable M B f"
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   326
proof%unimportant-
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   327
  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
   328
    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
   329
    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
   330
  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
   331
    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
   332
  ultimately show ?thesis 
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   333
    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
   334
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   335
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   336
lemma set_integral_Un_AE:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   337
  fixes f :: "_ \<Rightarrow> _ :: {banach, second_countable_topology}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   338
  assumes 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
   339
  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
   340
  and "set_integrable M B f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   341
  shows "LINT x:A\<union>B|M. f x = (LINT x:A|M. f x) + (LINT x:B|M. f x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   342
proof -
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   343
  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
   344
    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
   345
  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
   346
    using integrable_iff_bounded set_borel_measurable_def set_integrable_def by blast
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   347
  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
   348
  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
   349
    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
   350
      using ae by auto
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   351
    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
   352
      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
   353
  qed fact
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   354
  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
   355
    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
   356
  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
   357
    using ae
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 67339
diff changeset
   358
    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
   359
       (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
   360
  finally show ?thesis .
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   361
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   362
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   363
lemma%important set_integral_finite_Union:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   364
  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
   365
  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
   366
    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
   367
  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
   368
  using assms
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   369
proof%unimportant induction
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   370
  case (insert x F)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69173
diff changeset
   371
  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
   372
    by (meson disjoint_family_on_insert)
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   373
  with insert show ?case
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   374
    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
   375
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
   376
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   377
(* TODO: find a better name? *)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   378
lemma pos_integrable_to_top:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   379
  fixes l::real
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   380
  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
   381
  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
   382
  and intgbl: "\<And>i::nat. set_integrable M (A i) f"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   383
  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
   384
shows "set_integrable M (\<Union>i. A i) f"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   385
    unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   386
  apply (rule integrable_monotone_convergence[where f = "\<lambda>i::nat. \<lambda>x. indicator (A i) x *\<^sub>R f x" and x = l])
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   387
  apply (rule intgbl [unfolded set_integrable_def])
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   388
  prefer 3 apply (rule lim [unfolded set_lebesgue_integral_def])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   389
  apply (rule AE_I2)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60615
diff changeset
   390
  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
   391
proof (rule AE_I2)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   392
  { fix x assume "x \<in> space M"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   393
    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
   394
    proof cases
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   395
      assume "\<exists>i. x \<in> A i"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   396
      then guess i ..
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   397
      then have *: "eventually (\<lambda>i. x \<in> A i) sequentially"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 60615
diff changeset
   398
        using \<open>x \<in> A i\<close> \<open>mono A\<close> by (auto simp: eventually_sequentially mono_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   399
      show ?thesis
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   400
        apply (intro Lim_eventually)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   401
        using *
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   402
        apply eventually_elim
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   403
        apply (auto split: split_indicator)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   404
        done
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   405
    qed auto }
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   406
  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
   407
    apply (rule borel_measurable_LIMSEQ_real)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   408
    apply assumption
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   409
    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
   410
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   411
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   412
(* Proof from Royden Real Analysis, p. 91. *)
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   413
lemma%important lebesgue_integral_countable_add:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   414
  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
   415
  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
   416
    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
   417
    and intgbl: "set_integrable M (\<Union>i. A i) f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   418
  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
   419
    unfolding set_lebesgue_integral_def
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   420
proof%unimportant (subst integral_suminf[symmetric])
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   421
  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
   422
    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
   423
    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
   424
  { 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
   425
    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
   426
      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
   427
  note sums = this
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   428
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   429
  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
   430
    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
   431
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   432
  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
   433
    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
   434
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   435
  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
   436
  proof (rule summableI_nonneg_bounded)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   437
    fix n
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   438
    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
   439
      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
   440
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   441
    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
   442
      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
   443
    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
   444
      using norm_f
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   445
      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
   446
    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
   447
      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
   448
      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
   449
      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
   450
          apply (auto split: split_indicator)
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   451
      done
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   452
    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
   453
      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
   454
      by simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   455
  qed
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69173
diff changeset
   456
  show "LINT x|M. indicator (\<Union>(A ` UNIV)) x *\<^sub>R f x = LINT x|M. (\<Sum>i. indicator (A i) x *\<^sub>R f x)"
63886
685fb01256af move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
hoelzl
parents: 63627
diff changeset
   457
    apply (rule Bochner_Integration.integral_cong[OF refl])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   458
    apply (subst suminf_scaleR_left[OF sums_summable[OF indicator_sums, OF disj], symmetric])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   459
    using sums_unique[OF indicator_sums[OF disj]]
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   460
    apply auto
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   461
    done
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   462
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   463
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   464
lemma set_integral_cont_up:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   465
  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
   466
  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
   467
  and intgbl: "set_integrable M (\<Union>i. A i) f"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   468
shows "(\<lambda>i. LINT x:(A i)|M. f x) \<longlonglongrightarrow> LINT x:(\<Union>i. A i)|M. f x"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   469
  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
   470
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
   471
  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
   472
    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
   473
  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
   474
    using int_A integrable_iff_bounded set_integrable_def by blast
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69173
diff changeset
   475
  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
   476
    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
   477
  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
   478
    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
   479
    by fastforce
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   480
  { fix x i assume "x \<in> A i"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   481
    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
   482
      by (intro filterlim_cong refl)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   483
         (fastforce simp: eventually_sequentially incseq_def subset_eq intro!: exI[of _ i]) }
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   484
  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
   485
    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
   486
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
   487
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   488
(* 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
   489
lemma set_integral_cont_down:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   490
  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
   491
  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
   492
  and int0: "set_integrable M (A 0) f"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   493
  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
   494
  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
   495
proof (rule integral_dominated_convergence)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   496
  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
   497
    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
   498
  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
   499
    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
   500
  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
   501
    by force
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   502
  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
   503
    using int0 by (rule set_integrable_subset) (insert A, auto simp: decseq_def)
69313
b021008c5397 removed legacy input syntax
haftmann
parents: 69173
diff changeset
   504
  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
   505
                  "\<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
   506
    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
   507
  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
   508
    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
   509
  { fix x i assume "x \<in> space M" "x \<notin> A i"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   510
    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
   511
      by (intro filterlim_cong refl)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   512
         (auto split: split_indicator simp: eventually_sequentially decseq_def intro!: exI[of _ i]) }
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   513
  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
   514
    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
   515
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   516
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   517
lemma set_integral_at_point:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   518
  fixes a :: real
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   519
  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
   520
  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
   521
  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
   522
proof-
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   523
  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
   524
    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
   525
  then show ?thesis using assms
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   526
    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
   527
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   528
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   529
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   530
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
   531
  "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
   532
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   533
abbreviation complex_lebesgue_integral :: "'a measure \<Rightarrow> ('a \<Rightarrow> complex) \<Rightarrow> complex" ("integral\<^sup>C") where
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   534
  "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
   535
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   536
syntax
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   537
  "_complex_lebesgue_integral" :: "pttrn \<Rightarrow> complex \<Rightarrow> 'a measure \<Rightarrow> complex"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   538
 ("\<integral>\<^sup>C _. _ \<partial>_" [60,61] 110)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   539
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   540
translations
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   541
  "\<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
   542
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   543
syntax
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   544
  "_ascii_complex_lebesgue_integral" :: "pttrn \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   545
  ("(3CLINT _|_. _)" [0,110,60] 60)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   546
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   547
translations
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   548
  "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
   549
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   550
lemma complex_integrable_cnj [simp]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   551
  "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
   552
proof
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   553
  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
   554
  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
   555
    by (rule integrable_cnj)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   556
  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
   557
    by simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   558
qed simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   559
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   560
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
   561
  "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
   562
proof
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   563
  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
   564
  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
   565
    by (rule integrable_Re)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   566
  then show "integrable M f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   567
    by simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   568
qed simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   569
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   570
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   571
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
   572
  "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
   573
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   574
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
   575
  "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
   576
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   577
syntax
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   578
"_ascii_complex_set_lebesgue_integral" :: "pttrn \<Rightarrow> 'a set \<Rightarrow> 'a measure \<Rightarrow> real \<Rightarrow> real"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   579
("(4CLINT _:_|_. _)" [0,60,110,61] 60)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   580
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   581
translations
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   582
"CLINT x:A|M. f" == "CONST complex_set_lebesgue_integral M A (\<lambda>x. f)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   583
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   584
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
   585
  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
   586
  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
   587
  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
   588
  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
   589
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   590
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64284
diff changeset
   591
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
   592
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
   593
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   594
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
   595
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   596
syntax
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   597
"_set_nn_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   598
("(\<integral>\<^sup>+((_)\<in>(_)./ _)/\<partial>_)" [0,60,110,61] 60)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   599
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   600
"_set_lebesgue_integral" :: "pttrn => 'a set => 'a measure => ereal => ereal"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   601
("(\<integral>((_)\<in>(_)./ _)/\<partial>_)" [0,60,110,61] 60)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   602
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   603
translations
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   604
"\<integral>\<^sup>+x \<in> A. f \<partial>M" == "CONST set_nn_integral M A (\<lambda>x. f)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   605
"\<integral>x \<in> A. f \<partial>M" == "CONST set_lebesgue_integral M A (\<lambda>x. f)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   606
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   607
lemma nn_integral_disjoint_pair:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   608
  assumes [measurable]: "f \<in> borel_measurable M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   609
          "B \<in> sets M" "C \<in> sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   610
          "B \<inter> C = {}"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   611
  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
   612
proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   613
  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
   614
  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
   615
  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
   616
    by (auto split: split_indicator)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   617
  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
   618
    by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   619
  also have "... = (\<integral>\<^sup>+x. f x * indicator B x \<partial>M) + (\<integral>\<^sup>+x. f x * indicator C x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   620
    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
   621
  finally show ?thesis by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   622
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   623
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   624
lemma nn_integral_disjoint_pair_countspace:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   625
  assumes "B \<inter> C = {}"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   626
  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
   627
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
   628
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   629
lemma nn_integral_null_delta:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   630
  assumes "A \<in> sets M" "B \<in> sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   631
          "(A - B) \<union> (B - A) \<in> null_sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   632
  shows "(\<integral>\<^sup>+x \<in> A. f x \<partial>M) = (\<integral>\<^sup>+x \<in> B. f x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   633
proof (rule nn_integral_cong_AE, auto simp add: indicator_def)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   634
  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
   635
    using assms(3) AE_not_in by blast
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   636
  then show "AE a in M. a \<notin> A \<longrightarrow> a \<in> B \<longrightarrow> f a = 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   637
    by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   638
  show "AE x\<in>A in M. x \<notin> B \<longrightarrow> f x = 0"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   639
    using * by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   640
qed
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
lemma nn_integral_disjoint_family:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   643
  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
   644
      and "disjoint_family B"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   645
  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
   646
proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   647
  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
   648
    by (rule nn_integral_suminf) simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   649
  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
   650
  proof (cases)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   651
    assume "x \<in> (\<Union>n. B n)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   652
    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
   653
    have a: "finite {n}" by simp
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64284
diff changeset
   654
    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
   655
      by (metis IntI UNIV_I empty_iff)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   656
    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
   657
    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
   658
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   659
    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
   660
    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
   661
    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
   662
      by (metis sums_unique[OF sums_finite[OF a]])
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   663
    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
   664
    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
   665
    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
   666
    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
   667
  next
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   668
    assume "x \<notin> (\<Union>n. B n)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   669
    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
   670
    have "(\<Sum>n. f x * indicator (B n) x) = 0"
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64284
diff changeset
   671
      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
   672
    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
   673
  qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   674
  ultimately show ?thesis by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   675
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   676
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   677
lemma nn_set_integral_add:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   678
  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
   679
          "A \<in> sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   680
  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
   681
proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   682
  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
   683
    by (auto simp add: indicator_def intro!: nn_integral_cong)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   684
  also have "... = (\<integral>\<^sup>+x. f x * indicator A x \<partial>M) + (\<integral>\<^sup>+x. g x * indicator A x \<partial>M)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   685
    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
   686
  finally show ?thesis by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   687
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   688
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   689
lemma nn_set_integral_cong:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   690
  assumes "AE x in M. f x = g x"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   691
  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
   692
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
   693
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   694
lemma nn_set_integral_set_mono:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   695
  "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
   696
by (auto intro!: nn_integral_mono split: split_indicator)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   697
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   698
lemma nn_set_integral_mono:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   699
  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
   700
          "A \<in> sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   701
      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
   702
  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
   703
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
   704
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   705
lemma nn_set_integral_space [simp]:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   706
  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
   707
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
   708
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   709
lemma nn_integral_count_compose_inj:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   710
  assumes "inj_on g A"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   711
  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
   712
proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   713
  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
   714
    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
   715
  also have "... = (\<integral>\<^sup>+y. f y \<partial>count_space (g`A))"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   716
    by (simp add: assms nn_integral_bij_count_space inj_on_imp_bij_betw)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   717
  also have "... = (\<integral>\<^sup>+y \<in> g`A. f y \<partial>count_space UNIV)"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   718
    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
   719
  finally show ?thesis by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   720
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   721
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   722
lemma nn_integral_count_compose_bij:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   723
  assumes "bij_betw g A B"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   724
  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
   725
proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   726
  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
   727
  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
   728
    by (rule nn_integral_count_compose_inj)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   729
  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
   730
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   731
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   732
lemma set_integral_null_delta:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   733
  fixes f::"_ \<Rightarrow> _ :: {banach, second_countable_topology}"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   734
  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
   735
    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
   736
  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
   737
proof (rule set_integral_cong_set)
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   738
  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
   739
    using null AE_not_in by blast
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   740
  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
   741
    by auto
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   742
qed (simp_all add: set_borel_measurable_def)
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   743
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   744
lemma set_integral_space:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   745
  assumes "integrable M f"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   746
  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
   747
  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
   748
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   749
lemma null_if_pos_func_has_zero_nn_int:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   750
  fixes f::"'a \<Rightarrow> ennreal"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   751
  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
   752
    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
   753
  shows "A \<in> null_sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   754
proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   755
  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
   756
    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
   757
  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
   758
  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
   759
qed
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 null_if_pos_func_has_zero_int:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   762
  assumes [measurable]: "integrable M f" "A \<in> sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   763
      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
   764
  shows "A \<in> null_sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   765
proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   766
  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
   767
    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
   768
    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
   769
    by (auto simp: set_lebesgue_integral_def)
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   770
  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
   771
  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
   772
  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
   773
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   774
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64284
diff changeset
   775
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
   776
for nonnegative set integrals introduced earlier.\<close>
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   777
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   778
lemma (in sigma_finite_measure) density_unique2:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   779
  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
   780
  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
   781
  shows "AE x in M. f x = f' x"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   782
proof (rule density_unique)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   783
  show "density M f = density M f'"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   784
    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
   785
qed (auto simp add: assms)
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   786
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   787
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 64284
diff changeset
   788
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
   789
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
   790
only formulate it for real-valued functions.\<close>
64283
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   791
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   792
lemma density_unique_real:
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   793
  fixes f f'::"_ \<Rightarrow> real"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 67399
diff changeset
   794
  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
   795
  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
   796
  shows "AE x in M. f x = f' x"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   797
proof -
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   798
  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
   799
  then have [measurable]: "A \<in> sets M" by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   800
  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
   801
    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
   802
  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
   803
  then have "A \<in> null_sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   804
    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
   805
  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
   806
  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
   807
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   808
  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
   809
  then have [measurable]: "B \<in> sets M" by simp
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   810
  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
   811
    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
   812
  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
   813
  then have "B \<in> null_sets M"
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   814
    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
   815
  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
   816
  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
   817
  then show ?thesis using * by auto
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   818
qed
979cdfdf7a79 HOL-Probability: move conditional expectation from AFP/Ergodic_Theory
hoelzl
parents: 63958
diff changeset
   819
69566
c41954ee87cf more antiquotations -- less LaTeX macros;
wenzelm
parents: 69313
diff changeset
   820
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
   821
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
   822
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
   823
variations) are known as Scheffe lemma.
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   824
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   825
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
   826
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
   827
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   828
lemma%important Scheffe_lemma1:
64284
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   829
  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
   830
          "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
   831
          "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
   832
  shows "(\<lambda>n. \<integral>\<^sup>+ x. norm(F n x - f x) \<partial>M) \<longlonglongrightarrow> 0"
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   833
proof%unimportant -
64284
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   834
  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
   835
    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
   836
  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
   837
  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
   838
  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
   839
    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
   840
  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
   841
    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
   842
    by simp
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   843
  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
   844
    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
   845
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   846
  {
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   847
    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
   848
    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
   849
    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
   850
    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
   851
    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
   852
      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
   853
    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
   854
    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
   855
    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
   856
    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
   857
      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
   858
  }
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   859
  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
   860
  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
   861
    by (simp add: nn_integral_cong_AE ennreal_mult)
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   862
  also have "... = 2 * (\<integral>\<^sup>+ x. norm(f x) \<partial>M)" by (rule nn_integral_cmult) auto
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   863
  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
   864
    by simp
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   865
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   866
  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
   867
    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
   868
  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
   869
      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
   870
    by simp
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   871
  also have "... = (\<integral>\<^sup>+x. norm(f x) \<partial>M) + limsup (\<lambda>n. (\<integral>\<^sup>+x. norm(F n x) \<partial>M))"
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   872
    by (rule Limsup_const_add, auto simp add: finint)
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   873
  also have "... \<le> (\<integral>\<^sup>+x. norm(f x) \<partial>M) + (\<integral>\<^sup>+x. norm(f x) \<partial>M)"
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   874
    using assms(4) by (simp add: add_left_mono)
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   875
  also have "... = 2 * (\<integral>\<^sup>+x. norm(f x) \<partial>M)"
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   876
    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
   877
  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
   878
    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
   879
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   880
  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
   881
    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
   882
  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
   883
    by (rule nn_integral_mono)
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   884
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   885
  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
   886
    by (simp add: int_liminf)
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   887
  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
   888
    by (rule nn_integral_liminf) auto
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   889
  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
   890
    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
   891
  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
   892
    fix n
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   893
    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
   894
      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
   895
    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
   896
            = (\<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
   897
    proof (rule nn_integral_diff)
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   898
      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
   899
        by simp
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   900
      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
   901
        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
   902
      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
   903
    qed (auto simp add: assms)
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   904
    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
   905
      by simp
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   906
  qed
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   907
  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
   908
    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
   909
    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
   910
    by (intro add_mono) auto
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   911
  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
   912
    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
   913
    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
   914
  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
   915
    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
   916
  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
   917
    by fact
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   918
  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
   919
    using fin2 by simp
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   920
  then show ?thesis
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   921
    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
   922
qed
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   923
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   924
lemma%important Scheffe_lemma2:
64284
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   925
  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
   926
  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
   927
          "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
   928
          "\<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
   929
  shows "(\<lambda>n. \<integral>\<^sup>+ x. norm(F n x - f x) \<partial>M) \<longlonglongrightarrow> 0"
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   930
proof%unimportant (rule Scheffe_lemma1)
64284
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   931
  fix n::nat
f3b905b2eee2 HOL-Analysis: more theorems from Sébastien Gouëzel's Ergodic_Theory
hoelzl
parents: 64283
diff changeset
   932
  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
   933
  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
   934
  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
   935
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
   936
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   937
lemma%important tendsto_set_lebesgue_integral_at_right:
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   938
  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
   939
  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
   940
      and "set_integrable M {a<..b} f"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   941
  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
   942
             set_lebesgue_integral M {a<..b} f) (at_right a)"
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   943
proof%unimportant (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
   944
  case (1 S)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   945
  have eq: "(\<Union>n. {S n..b}) = {a<..b}"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   946
  proof safe
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   947
    fix x n assume "x \<in> {S n..b}"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   948
    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
   949
  next
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   950
    fix x assume "x \<in> {a<..b}"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   951
    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
   952
      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
   953
  qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   954
  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
   955
    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
   956
  with eq show ?case by simp
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   957
qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   958
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   959
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   960
text \<open>
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   961
  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
   962
  improper integrals.
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   963
\<close>
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   964
lemma%important tendsto_set_lebesgue_integral_at_left:
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   965
  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
   966
  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
   967
      and "set_integrable M {a..<b} f"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   968
  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
   969
             set_lebesgue_integral M {a..<b} f) (at_left b)"
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   970
proof%unimportant (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
   971
  case (1 S)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   972
  have eq: "(\<Union>n. {a..S n}) = {a..<b}"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   973
  proof safe
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   974
    fix x n assume "x \<in> {a..S n}"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   975
    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
   976
  next
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   977
    fix x assume "x \<in> {a..<b}"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   978
    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
   979
      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
   980
  qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   981
  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
   982
    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
   983
  with eq show ?case by simp
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   984
qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   985
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   986
lemma%important tendsto_set_lebesgue_integral_at_top:
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   987
  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   988
  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
   989
      and int: "set_integrable M {a..} f"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   990
  shows "((\<lambda>b. set_lebesgue_integral M {a..b} f) \<longlongrightarrow> set_lebesgue_integral M {a..} f) at_top"
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
   991
proof%unimportant (rule tendsto_at_topI_sequentially)
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   992
  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
   993
  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
   994
    unfolding set_lebesgue_integral_def
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   995
  proof (rule integral_dominated_convergence)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
   996
    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
   997
      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
   998
    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
   999
    proof
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1000
      fix x
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1001
      from \<open>filterlim X at_top sequentially\<close>
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1002
      have "eventually (\<lambda>n. x \<le> X n) sequentially"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1003
        unfolding filterlim_at_top_ge[where c=x] by auto
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1004
      then show "(\<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
  1005
        by (intro Lim_eventually) (auto split: split_indicator elim!: eventually_mono)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1006
    qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1007
    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
  1008
                             indicator {a..} x *\<^sub>R norm (f x)"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1009
      by (auto split: split_indicator)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1010
  next
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1011
    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
  1012
      by (simp add: set_integrable_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1013
  next
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1014
    fix n :: nat
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1015
    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
  1016
    with int have "set_integrable M {a..X n} f"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1017
      by (rule set_integrable_subset) auto
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1018
    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
  1019
      by (simp add: set_integrable_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1020
  qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1021
qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1022
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
  1023
lemma%important tendsto_set_lebesgue_integral_at_bot:
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1024
  fixes f :: "real \<Rightarrow> 'a::{banach, second_countable_topology}"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1025
  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
  1026
      and int: "set_integrable M {..b} f"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1027
    shows "((\<lambda>a. set_lebesgue_integral M {a..b} f) \<longlongrightarrow> set_lebesgue_integral M {..b} f) at_bot"
69173
38beaaebe736 tagged 8 theories for the Analysis manual.
Angeliki KoutsoukouArgyraki <ak2110@cam.ac.uk>
parents: 68721
diff changeset
  1028
proof%unimportant (rule tendsto_at_botI_sequentially)
68721
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1029
  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
  1030
  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
  1031
    unfolding set_lebesgue_integral_def
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1032
  proof (rule integral_dominated_convergence)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1033
    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
  1034
      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
  1035
    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
  1036
    proof
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1037
      fix x
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1038
      from \<open>filterlim X at_bot sequentially\<close>
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1039
      have "eventually (\<lambda>n. x \<ge> X n) sequentially"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1040
        unfolding filterlim_at_bot_le[where c=x] by auto
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1041
      then show "(\<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
  1042
        by (intro Lim_eventually) (auto split: split_indicator elim!: eventually_mono)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1043
    qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1044
    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
  1045
                             indicator {..b} x *\<^sub>R norm (f x)"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1046
      by (auto split: split_indicator)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1047
  next
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1048
    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
  1049
      by (simp add: set_integrable_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1050
  next
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1051
    fix n :: nat
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1052
    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
  1053
    with int have "set_integrable M {X n..b} f"
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1054
      by (rule set_integrable_subset) auto
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1055
    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
  1056
      by (simp add: set_integrable_def)
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1057
  qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1058
qed
53ad5c01be3f Small lemmas about analysis
eberlm <eberlm@in.tum.de>
parents: 68403
diff changeset
  1059
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1060
end