src/HOL/Analysis/Interval_Integral.thy
author paulson <lp15@cam.ac.uk>
Mon, 11 May 2020 11:15:41 +0100
changeset 71827 5e315defb038
parent 70613 8b7f6ecb3369
child 73526 a3cc9fa1295d
permissions -rw-r--r--
the Uniq quantifier
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/Interval_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)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     3
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     4
Lebesgue integral over an interval (with endpoints possibly +-\<infinity>)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     5
*)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     6
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
     7
theory Interval_Integral (*FIX ME rename? Lebesgue  *)
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
     8
  imports Equivalence_Lebesgue_Henstock_Integration
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
     9
begin
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    10
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    11
definition "einterval a b = {x. a < ereal x \<and> ereal x < b}"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    12
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    13
lemma einterval_eq[simp]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    14
  shows einterval_eq_Icc: "einterval (ereal a) (ereal b) = {a <..< b}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    15
    and einterval_eq_Ici: "einterval (ereal a) \<infinity> = {a <..}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    16
    and einterval_eq_Iic: "einterval (- \<infinity>) (ereal b) = {..< b}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    17
    and einterval_eq_UNIV: "einterval (- \<infinity>) \<infinity> = UNIV"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    18
  by (auto simp: einterval_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    19
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    20
lemma einterval_same: "einterval a a = {}"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
    21
  by (auto simp: einterval_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    22
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    23
lemma einterval_iff: "x \<in> einterval a b \<longleftrightarrow> a < ereal x \<and> ereal x < b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    24
  by (simp add: einterval_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    25
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    26
lemma einterval_nonempty: "a < b \<Longrightarrow> \<exists>c. c \<in> einterval a b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    27
  by (cases a b rule: ereal2_cases, auto simp: einterval_def intro!: dense gt_ex lt_ex)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    28
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    29
lemma open_einterval[simp]: "open (einterval a b)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    30
  by (cases a b rule: ereal2_cases)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    31
     (auto simp: einterval_def intro!: open_Collect_conj open_Collect_less continuous_intros)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    32
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    33
lemma borel_einterval[measurable]: "einterval a b \<in> sets borel"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    34
  unfolding einterval_def by measurable
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    35
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
    36
subsection \<open>Approximating a (possibly infinite) interval\<close>
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    37
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    38
lemma filterlim_sup1: "(LIM x F. f x :> G1) \<Longrightarrow> (LIM x F. f x :> (sup G1 G2))"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    39
 unfolding filterlim_def by (auto intro: le_supI1)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    40
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    41
lemma ereal_incseq_approx:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    42
  fixes a b :: ereal
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    43
  assumes "a < b"
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
    44
  obtains X :: "nat \<Rightarrow> real" where "incseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X \<longlonglongrightarrow> b"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    45
proof (cases b)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    46
  case PInf
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
    47
  with \<open>a < b\<close> have "a = -\<infinity> \<or> (\<exists>r. a = ereal r)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    48
    by (cases a) auto
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
    49
  moreover have "(\<lambda>x. ereal (real (Suc x))) \<longlonglongrightarrow> \<infinity>"
71827
5e315defb038 the Uniq quantifier
paulson <lp15@cam.ac.uk>
parents: 70613
diff changeset
    50
    by (simp add: Lim_PInfty filterlim_sequentially_Suc) (metis le_SucI of_nat_Suc of_nat_mono order_trans real_arch_simple)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
    51
  moreover have "\<And>r. (\<lambda>x. ereal (r + real (Suc x))) \<longlonglongrightarrow> \<infinity>"
71827
5e315defb038 the Uniq quantifier
paulson <lp15@cam.ac.uk>
parents: 70613
diff changeset
    52
    by (simp add: filterlim_sequentially_Suc Lim_PInfty) (metis add.commute diff_le_eq nat_ceiling_le_eq)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    53
  ultimately show thesis
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
    54
    by (intro that[of "\<lambda>i. real_of_ereal a + Suc i"])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    55
       (auto simp: incseq_def PInf)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    56
next
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    57
  case (real b')
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62390
diff changeset
    58
  define d where "d = b' - (if a = -\<infinity> then b' - 1 else real_of_ereal a)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
    59
  with \<open>a < b\<close> have a': "0 < d"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    60
    by (cases a) (auto simp: real)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    61
  moreover
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    62
  have "\<And>i r. r < b' \<Longrightarrow> (b' - r) * 1 < (b' - r) * real (Suc (Suc i))"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    63
    by (intro mult_strict_left_mono) auto
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
    64
  with \<open>a < b\<close> a' have "\<And>i. a < ereal (b' - d / real (Suc (Suc i)))"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    65
    by (cases a) (auto simp: real d_def field_simps)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
    66
  moreover
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
    67
  have "(\<lambda>i. b' - d / real i) \<longlonglongrightarrow> b'"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
    68
    by (force intro: tendsto_eq_intros tendsto_divide_0[OF tendsto_const] filterlim_sup1
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
    69
              simp: at_infinity_eq_at_top_bot filterlim_real_sequentially)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
    70
  then have "(\<lambda>i. b' - d / Suc (Suc i)) \<longlonglongrightarrow> b'"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
    71
    by (blast intro: dest: filterlim_sequentially_Suc [THEN iffD2])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    72
  ultimately show thesis
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    73
    by (intro that[of "\<lambda>i. b' - d / Suc (Suc i)"])
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
    74
       (auto simp: real incseq_def intro!: divide_left_mono)
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
    75
qed (insert \<open>a < b\<close>, auto)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    76
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    77
lemma ereal_decseq_approx:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    78
  fixes a b :: ereal
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    79
  assumes "a < b"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
    80
  obtains X :: "nat \<Rightarrow> real" where
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
    81
    "decseq X" "\<And>i. a < X i" "\<And>i. X i < b" "X \<longlonglongrightarrow> a"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    82
proof -
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
    83
  have "-b < -a" using \<open>a < b\<close> by simp
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    84
  from ereal_incseq_approx[OF this] guess X .
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    85
  then show thesis
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    86
    apply (intro that[of "\<lambda>i. - X i"])
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68096
diff changeset
    87
    apply (auto simp: decseq_def incseq_def simp flip: uminus_ereal.simps)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    88
    apply (metis ereal_minus_less_minus ereal_uminus_uminus ereal_Lim_uminus)+
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    89
    done
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    90
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    91
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    92
proposition einterval_Icc_approximation:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    93
  fixes a b :: ereal
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    94
  assumes "a < b"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
    95
  obtains u l :: "nat \<Rightarrow> real" where
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    96
    "einterval a b = (\<Union>i. {l i .. u i})"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
    97
    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
    98
    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
    99
proof -
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   100
  from dense[OF \<open>a < b\<close>] obtain c where "a < c" "c < b" by safe
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   101
  from ereal_incseq_approx[OF \<open>c < b\<close>] guess u . note u = this
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   102
  from ereal_decseq_approx[OF \<open>a < c\<close>] guess l . note l = this
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   103
  { fix i from less_trans[OF \<open>l i < c\<close> \<open>c < u i\<close>] have "l i < u i" by simp }
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   104
  have "einterval a b = (\<Union>i. {l i .. u i})"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   105
  proof (auto simp: einterval_iff)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   106
    fix x assume "a < ereal x" "ereal x < b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   107
    have "eventually (\<lambda>i. ereal (l i) < ereal x) sequentially"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   108
      using l(4) \<open>a < ereal x\<close> by (rule order_tendstoD)
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   109
    moreover
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   110
    have "eventually (\<lambda>i. ereal x < ereal (u i)) sequentially"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   111
      using u(4) \<open>ereal x< b\<close> by (rule order_tendstoD)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   112
    ultimately have "eventually (\<lambda>i. l i < x \<and> x < u i) sequentially"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   113
      by eventually_elim auto
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   114
    then show "\<exists>i. l i \<le> x \<and> x \<le> u i"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   115
      by (auto intro: less_imp_le simp: eventually_sequentially)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   116
  next
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   117
    fix x i assume "l i \<le> x" "x \<le> u i"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   118
    with \<open>a < ereal (l i)\<close> \<open>ereal (u i) < b\<close>
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   119
    show "a < ereal x" "ereal x < b"
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68096
diff changeset
   120
      by (auto simp flip: ereal_less_eq(3))
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   121
  qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   122
  show thesis
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   123
    by (intro that) fact+
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   124
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   125
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   126
(* TODO: in this definition, it would be more natural if einterval a b included a and b when
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   127
   they are real. *)
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69683
diff changeset
   128
definition\<^marker>\<open>tag important\<close> interval_lebesgue_integral :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a) \<Rightarrow> 'a::{banach, second_countable_topology}" where
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   129
  "interval_lebesgue_integral M a b f =
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   130
    (if a \<le> b then (LINT x:einterval a b|M. f x) else - (LINT x:einterval b a|M. f x))"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   131
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   132
syntax
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   133
  "_ascii_interval_lebesgue_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real measure \<Rightarrow> real \<Rightarrow> real"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   134
  ("(5LINT _=_.._|_. _)" [0,60,60,61,100] 60)
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
translations
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   137
  "LINT x=a..b|M. f" == "CONST interval_lebesgue_integral M a b (\<lambda>x. f)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   138
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69683
diff changeset
   139
definition\<^marker>\<open>tag important\<close> interval_lebesgue_integrable :: "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> 'a::{banach, second_countable_topology}) \<Rightarrow> bool" where
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   140
  "interval_lebesgue_integrable M a b f =
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   141
    (if a \<le> b then set_integrable M (einterval a b) f else set_integrable M (einterval b a) f)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   142
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   143
syntax
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   144
  "_ascii_interval_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real \<Rightarrow> real"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   145
  ("(4LBINT _=_.._. _)" [0,60,60,61] 60)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   146
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   147
translations
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   148
  "LBINT x=a..b. f" == "CONST interval_lebesgue_integral CONST lborel a b (\<lambda>x. f)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   149
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   150
subsection\<open>Basic properties of integration over an interval\<close>
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   151
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   152
lemma interval_lebesgue_integral_cong:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   153
  "a \<le> b \<Longrightarrow> (\<And>x. x \<in> einterval a b \<Longrightarrow> f x = g x) \<Longrightarrow> einterval a b \<in> sets M \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   154
    interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   155
  by (auto intro: set_lebesgue_integral_cong simp: interval_lebesgue_integral_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   156
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   157
lemma interval_lebesgue_integral_cong_AE:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   158
  "f \<in> borel_measurable M \<Longrightarrow> g \<in> borel_measurable M \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   159
    a \<le> b \<Longrightarrow> AE x \<in> einterval a b in M. f x = g x \<Longrightarrow> einterval a b \<in> sets M \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   160
    interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   161
  by (auto intro: set_lebesgue_integral_cong_AE simp: interval_lebesgue_integral_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   162
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   163
lemma interval_integrable_mirror:
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   164
  shows "interval_lebesgue_integrable lborel a b (\<lambda>x. f (-x)) \<longleftrightarrow>
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   165
    interval_lebesgue_integrable lborel (-b) (-a) f"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   166
proof -
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   167
  have *: "indicator (einterval a b) (- x) = (indicator (einterval (-b) (-a)) x :: real)"
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   168
    for a b :: ereal and x :: real
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   169
    by (cases a b rule: ereal2_cases) (auto simp: einterval_def split: split_indicator)
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   170
  show ?thesis
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   171
    unfolding interval_lebesgue_integrable_def
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   172
    using lborel_integrable_real_affine_iff[symmetric, of "-1" "\<lambda>x. indicator (einterval _ _) x *\<^sub>R f x" 0]
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   173
    by (simp add: * set_integrable_def)
62083
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   174
qed
7582b39f51ed add the proof of the central limit theorem
hoelzl
parents: 61973
diff changeset
   175
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   176
lemma interval_lebesgue_integral_add [intro, simp]:
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   177
  fixes M a b f
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   178
  assumes "interval_lebesgue_integrable M a b f" "interval_lebesgue_integrable M a b g"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   179
  shows "interval_lebesgue_integrable M a b (\<lambda>x. f x + g x)" and
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   180
    "interval_lebesgue_integral M a b (\<lambda>x. f x + g x) =
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   181
   interval_lebesgue_integral M a b f + interval_lebesgue_integral M a b g"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   182
using assms by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   183
    field_simps)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   184
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   185
lemma interval_lebesgue_integral_diff [intro, simp]:
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   186
  fixes M a b f
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   187
  assumes "interval_lebesgue_integrable M a b f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   188
    "interval_lebesgue_integrable M a b g"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   189
  shows "interval_lebesgue_integrable M a b (\<lambda>x. f x - g x)" and
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   190
    "interval_lebesgue_integral M a b (\<lambda>x. f x - g x) =
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   191
   interval_lebesgue_integral M a b f - interval_lebesgue_integral M a b g"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   192
using assms by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   193
    field_simps)
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 interval_lebesgue_integrable_mult_right [intro, simp]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   196
  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   197
  shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   198
    interval_lebesgue_integrable M a b (\<lambda>x. c * f x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   199
  by (simp add: interval_lebesgue_integrable_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   200
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   201
lemma interval_lebesgue_integrable_mult_left [intro, simp]:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   202
  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, 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
   203
  shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   204
    interval_lebesgue_integrable M a b (\<lambda>x. f x * c)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   205
  by (simp add: interval_lebesgue_integrable_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   206
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   207
lemma interval_lebesgue_integrable_divide [intro, simp]:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59587
diff changeset
   208
  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, 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
   209
  shows "(c \<noteq> 0 \<Longrightarrow> interval_lebesgue_integrable M a b f) \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   210
    interval_lebesgue_integrable M a b (\<lambda>x. f x / c)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   211
  by (simp add: interval_lebesgue_integrable_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   212
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   213
lemma interval_lebesgue_integral_mult_right [simp]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   214
  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, second_countable_topology}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   215
  shows "interval_lebesgue_integral M a b (\<lambda>x. c * f x) =
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   216
    c * interval_lebesgue_integral M a b f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   217
  by (simp add: interval_lebesgue_integral_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   218
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   219
lemma interval_lebesgue_integral_mult_left [simp]:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   220
  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, 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
   221
  shows "interval_lebesgue_integral M a b (\<lambda>x. f x * c) =
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   222
    interval_lebesgue_integral M a b f * c"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   223
  by (simp add: interval_lebesgue_integral_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   224
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   225
lemma interval_lebesgue_integral_divide [simp]:
59867
58043346ca64 given up separate type classes demanding `inverse 0 = 0`
haftmann
parents: 59587
diff changeset
   226
  fixes M a b c and f :: "real \<Rightarrow> 'a::{banach, real_normed_field, field, second_countable_topology}"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   227
  shows "interval_lebesgue_integral M a b (\<lambda>x. f x / c) =
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   228
    interval_lebesgue_integral M a b f / c"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   229
  by (simp add: interval_lebesgue_integral_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   230
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   231
lemma interval_lebesgue_integral_uminus:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   232
  "interval_lebesgue_integral M a b (\<lambda>x. - f x) = - interval_lebesgue_integral M a b f"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   233
  by (auto simp: interval_lebesgue_integral_def interval_lebesgue_integrable_def set_lebesgue_integral_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   234
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   235
lemma interval_lebesgue_integral_of_real:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   236
  "interval_lebesgue_integral M a b (\<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
   237
    of_real (interval_lebesgue_integral M a b f)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   238
  unfolding interval_lebesgue_integral_def
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   239
  by (auto simp: interval_lebesgue_integral_def set_integral_complex_of_real)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   240
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   241
lemma interval_lebesgue_integral_le_eq:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   242
  fixes a b f
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   243
  assumes "a \<le> b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   244
  shows "interval_lebesgue_integral M a b f = (LINT x : einterval a b | M. f x)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   245
  using assms by (auto simp: interval_lebesgue_integral_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   246
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   247
lemma interval_lebesgue_integral_gt_eq:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   248
  fixes a b f
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   249
  assumes "a > b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   250
  shows "interval_lebesgue_integral M a b f = -(LINT x : einterval b a | M. f x)"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   251
using assms by (auto simp: interval_lebesgue_integral_def less_imp_le einterval_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   252
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   253
lemma interval_lebesgue_integral_gt_eq':
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   254
  fixes a b f
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   255
  assumes "a > b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   256
  shows "interval_lebesgue_integral M a b f = - interval_lebesgue_integral M b a f"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   257
using assms by (auto simp: interval_lebesgue_integral_def less_imp_le einterval_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   258
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   259
lemma interval_integral_endpoints_same [simp]: "(LBINT x=a..a. f x) = 0"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   260
  by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def einterval_same)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   261
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   262
lemma interval_integral_endpoints_reverse: "(LBINT x=a..b. f x) = -(LBINT x=b..a. f x)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   263
  by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integral_def set_lebesgue_integral_def einterval_same)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   264
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   265
lemma interval_integrable_endpoints_reverse:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   266
  "interval_lebesgue_integrable lborel a b f \<longleftrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   267
    interval_lebesgue_integrable lborel b a f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   268
  by (cases a b rule: linorder_cases) (auto simp: interval_lebesgue_integrable_def einterval_same)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   269
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   270
lemma interval_integral_reflect:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   271
  "(LBINT x=a..b. f x) = (LBINT x=-b..-a. f (-x))"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   272
proof (induct a b rule: linorder_wlog)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   273
  case (sym a b) then show ?case
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   274
    by (auto simp: interval_lebesgue_integral_def interval_integrable_endpoints_reverse
62390
842917225d56 more canonical names
nipkow
parents: 62083
diff changeset
   275
             split: if_split_asm)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   276
next
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   277
  case (le a b) 
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   278
  have "LBINT x:{x. - x \<in> einterval a b}. f (- x) = LBINT x:einterval (- b) (- a). f (- x)"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   279
    unfolding interval_lebesgue_integrable_def set_lebesgue_integral_def
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   280
    apply (rule Bochner_Integration.integral_cong [OF refl])
68046
6aba668aea78 new simp modifier: reorient
nipkow
parents: 67974
diff changeset
   281
    by (auto simp: einterval_iff ereal_uminus_le_reorder ereal_uminus_less_reorder not_less
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68096
diff changeset
   282
             simp flip: uminus_ereal.simps
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   283
             split: split_indicator)
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   284
  then show ?case
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   285
    unfolding interval_lebesgue_integral_def 
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   286
    by (subst set_integral_reflect) (simp add: le)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   287
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   288
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   289
lemma interval_lebesgue_integral_0_infty:
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   290
  "interval_lebesgue_integrable M 0 \<infinity> f \<longleftrightarrow> set_integrable M {0<..} f"
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   291
  "interval_lebesgue_integral M 0 \<infinity> f = (LINT x:{0<..}|M. f x)"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   292
  unfolding zero_ereal_def
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   293
  by (auto simp: interval_lebesgue_integral_le_eq interval_lebesgue_integrable_def)
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   294
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   295
lemma interval_integral_to_infinity_eq: "(LINT x=ereal a..\<infinity> | M. f x) = (LINT x : {a<..} | M. f x)"
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   296
  unfolding interval_lebesgue_integral_def by auto
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   297
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   298
proposition interval_integrable_to_infinity_eq: "(interval_lebesgue_integrable M a \<infinity> f) =
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   299
  (set_integrable M {a<..} f)"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69683
diff changeset
   300
  unfolding interval_lebesgue_integrable_def by auto
61897
bc0fc5499085 Bochner integral: prove dominated convergence at_top
hoelzl
parents: 61882
diff changeset
   301
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   302
subsection\<open>Basic properties of integration over an interval wrt lebesgue measure\<close>
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   303
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   304
lemma interval_integral_zero [simp]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   305
  fixes a b :: ereal
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   306
  shows "LBINT x=a..b. 0 = 0"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   307
unfolding interval_lebesgue_integral_def set_lebesgue_integral_def einterval_eq
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   308
by simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   309
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   310
lemma interval_integral_const [intro, simp]:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   311
  fixes a b c :: real
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   312
  shows "interval_lebesgue_integrable lborel a b (\<lambda>x. c)" and "LBINT x=a..b. c = c * (b - a)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   313
  unfolding interval_lebesgue_integral_def interval_lebesgue_integrable_def einterval_eq
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   314
  by (auto simp: less_imp_le field_simps measure_def 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
   315
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   316
lemma interval_integral_cong_AE:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   317
  assumes [measurable]: "f \<in> borel_measurable borel" "g \<in> borel_measurable borel"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   318
  assumes "AE x \<in> einterval (min a b) (max a b) in lborel. f x = g x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   319
  shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   320
  using assms
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   321
proof (induct a b rule: linorder_wlog)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   322
  case (sym a b) then show ?case
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   323
    by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   324
next
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   325
  case (le a b) then show ?case
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   326
    by (auto simp: interval_lebesgue_integral_def max_def min_def
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   327
             intro!: set_lebesgue_integral_cong_AE)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   328
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   329
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   330
lemma interval_integral_cong:
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   331
  assumes "\<And>x. x \<in> einterval (min a b) (max a b) \<Longrightarrow> f x = g x"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   332
  shows "interval_lebesgue_integral lborel a b f = interval_lebesgue_integral lborel a b g"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   333
  using assms
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   334
proof (induct a b rule: linorder_wlog)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   335
  case (sym a b) then show ?case
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   336
    by (simp add: min.commute max.commute interval_integral_endpoints_reverse[of a b])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   337
next
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   338
  case (le a b) then show ?case
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   339
    by (auto simp: interval_lebesgue_integral_def max_def min_def
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   340
             intro!: set_lebesgue_integral_cong)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   341
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   342
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   343
lemma interval_lebesgue_integrable_cong_AE:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   344
    "f \<in> borel_measurable lborel \<Longrightarrow> g \<in> borel_measurable lborel \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   345
    AE x \<in> einterval (min a b) (max a b) in lborel. f x = g x \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   346
    interval_lebesgue_integrable lborel a b f = interval_lebesgue_integrable lborel a b g"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   347
  apply (simp add: interval_lebesgue_integrable_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   348
  apply (intro conjI impI set_integrable_cong_AE)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   349
  apply (auto simp: min_def max_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   350
  done
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   351
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   352
lemma interval_integrable_abs_iff:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   353
  fixes f :: "real \<Rightarrow> real"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   354
  shows  "f \<in> borel_measurable lborel \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   355
    interval_lebesgue_integrable lborel a b (\<lambda>x. \<bar>f x\<bar>) = interval_lebesgue_integrable lborel a b f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   356
  unfolding interval_lebesgue_integrable_def
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   357
  by (subst (1 2) set_integrable_abs_iff') simp_all
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   358
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   359
lemma interval_integral_Icc:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   360
  fixes a b :: real
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   361
  shows "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {a..b}. f x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   362
  by (auto intro!: set_integral_discrete_difference[where X="{a, b}"]
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   363
           simp add: interval_lebesgue_integral_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   364
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   365
lemma interval_integral_Icc':
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   366
  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {x. a \<le> ereal x \<and> ereal x \<le> b}. f x)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   367
  by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"]
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   368
           simp add: interval_lebesgue_integral_def einterval_iff)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   369
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   370
lemma interval_integral_Ioc:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   371
  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {a<..b}. f x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   372
  by (auto intro!: set_integral_discrete_difference[where X="{a, b}"]
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   373
           simp add: interval_lebesgue_integral_def einterval_iff)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   374
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   375
(* TODO: other versions as well? *) (* Yes: I need the Icc' version. *)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   376
lemma interval_integral_Ioc':
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   377
  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {x. a < ereal x \<and> ereal x \<le> b}. f x)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   378
  by (auto intro!: set_integral_discrete_difference[where X="{real_of_ereal a, real_of_ereal b}"]
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   379
           simp add: interval_lebesgue_integral_def einterval_iff)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   380
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   381
lemma interval_integral_Ico:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   382
  "a \<le> b \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {a..<b}. f x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   383
  by (auto intro!: set_integral_discrete_difference[where X="{a, b}"]
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   384
           simp add: interval_lebesgue_integral_def einterval_iff)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   385
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   386
lemma interval_integral_Ioi:
61882
8b4b5d74c587 Probability: fix coercions (real ~> real_of_enat)
hoelzl
parents: 61808
diff changeset
   387
  "\<bar>a\<bar> < \<infinity> \<Longrightarrow> (LBINT x=a..\<infinity>. f x) = (LBINT x : {real_of_ereal a <..}. f x)"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   388
  by (auto simp: interval_lebesgue_integral_def einterval_iff)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   389
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   390
lemma interval_integral_Ioo:
61882
8b4b5d74c587 Probability: fix coercions (real ~> real_of_enat)
hoelzl
parents: 61808
diff changeset
   391
  "a \<le> b \<Longrightarrow> \<bar>a\<bar> < \<infinity> ==> \<bar>b\<bar> < \<infinity> \<Longrightarrow> (LBINT x=a..b. f x) = (LBINT x : {real_of_ereal a <..< real_of_ereal b}. f x)"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   392
  by (auto simp: interval_lebesgue_integral_def einterval_iff)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   393
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   394
lemma interval_integral_discrete_difference:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   395
  fixes f :: "real \<Rightarrow> 'b::{banach, second_countable_topology}" and a b :: ereal
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   396
  assumes "countable X"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   397
  and eq: "\<And>x. a \<le> b \<Longrightarrow> a < x \<Longrightarrow> x < b \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   398
  and anti_eq: "\<And>x. b \<le> a \<Longrightarrow> b < x \<Longrightarrow> x < a \<Longrightarrow> x \<notin> X \<Longrightarrow> f x = g x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   399
  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
   400
  shows "interval_lebesgue_integral M a b f = interval_lebesgue_integral M a b g"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   401
  unfolding interval_lebesgue_integral_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
   402
  apply (intro if_cong refl arg_cong[where f="\<lambda>x. - x"] integral_discrete_difference[of X] assms)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   403
  apply (auto simp: eq anti_eq einterval_iff 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
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   406
lemma interval_integral_sum:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   407
  fixes a b c :: ereal
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   408
  assumes integrable: "interval_lebesgue_integrable lborel (min a (min b c)) (max a (max b c)) f"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   409
  shows "(LBINT x=a..b. f x) + (LBINT x=b..c. f x) = (LBINT x=a..c. f x)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   410
proof -
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   411
  let ?I = "\<lambda>a b. LBINT x=a..b. f x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   412
  { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \<le> b" "b \<le> c"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   413
    then have ord: "a \<le> b" "b \<le> c" "a \<le> c" and f': "set_integrable lborel (einterval a c) f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   414
      by (auto simp: interval_lebesgue_integrable_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   415
    then have f: "set_borel_measurable borel (einterval a c) f"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   416
      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
   417
      by (drule_tac borel_measurable_integrable) simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   418
    have "(LBINT x:einterval a c. f x) = (LBINT x:einterval a b \<union> einterval b c. f x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   419
    proof (rule set_integral_cong_set)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   420
      show "AE x in lborel. (x \<in> einterval a b \<union> einterval b c) = (x \<in> einterval a c)"
61609
77b453bd616f Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
paulson <lp15@cam.ac.uk>
parents: 59867
diff changeset
   421
        using AE_lborel_singleton[of "real_of_ereal b"] ord
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   422
        by (cases a b c rule: ereal3_cases) (auto simp: einterval_iff)
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   423
      show "set_borel_measurable lborel (einterval a c) f" "set_borel_measurable lborel (einterval a b \<union> einterval b c) f"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   424
        unfolding set_borel_measurable_def
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   425
        using ord by (auto simp: einterval_iff intro!: set_borel_measurable_subset[OF f, unfolded set_borel_measurable_def])
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   426
    qed
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   427
    also have "\<dots> = (LBINT x:einterval a b. f x) + (LBINT x:einterval b c. f x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   428
      using ord
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   429
      by (intro set_integral_Un_AE) (auto intro!: set_integrable_subset[OF f'] simp: einterval_iff not_less)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   430
    finally have "?I a b + ?I b c = ?I a c"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   431
      using ord by (simp add: interval_lebesgue_integral_def)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   432
  } note 1 = this
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   433
  { fix a b c :: ereal assume "interval_lebesgue_integrable lborel a c f" "a \<le> b" "b \<le> c"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   434
    from 1[OF this] have "?I b c + ?I a b = ?I a c"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   435
      by (metis add.commute)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   436
  } note 2 = this
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   437
  have 3: "\<And>a b. b \<le> a \<Longrightarrow> (LBINT x=a..b. f x) = - (LBINT x=b..a. f x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   438
    by (rule interval_integral_endpoints_reverse)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   439
  show ?thesis
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   440
    using integrable
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   441
    by (cases a b b c a c rule: linorder_le_cases[case_product linorder_le_cases linorder_cases])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   442
       (simp_all add: min_absorb1 min_absorb2 max_absorb1 max_absorb2 field_simps 1 2 3)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   443
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   444
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   445
lemma interval_integrable_isCont:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   446
  fixes a b and f :: "real \<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
   447
  shows "(\<And>x. min a b \<le> x \<Longrightarrow> x \<le> max a b \<Longrightarrow> isCont f x) \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   448
    interval_lebesgue_integrable lborel a b f"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   449
proof (induct a b rule: linorder_wlog)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   450
  case (le a b) then show ?case
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   451
    unfolding interval_lebesgue_integrable_def set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   452
    by (auto simp: interval_lebesgue_integrable_def
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   453
        intro!: set_integrable_subset[unfolded set_integrable_def, OF borel_integrable_compact[of "{a .. b}"]]
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   454
        continuous_at_imp_continuous_on)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   455
qed (auto intro: interval_integrable_endpoints_reverse[THEN iffD1])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   456
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   457
lemma interval_integrable_continuous_on:
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   458
  fixes a b :: real and f
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   459
  assumes "a \<le> b" and "continuous_on {a..b} f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   460
  shows "interval_lebesgue_integrable lborel a b f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   461
using assms unfolding interval_lebesgue_integrable_def apply simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   462
  by (rule set_integrable_subset, rule borel_integrable_atLeastAtMost' [of a b], auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   463
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   464
lemma interval_integral_eq_integral:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   465
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   466
  shows "a \<le> b \<Longrightarrow> set_integrable lborel {a..b} f \<Longrightarrow> LBINT x=a..b. f x = integral {a..b} f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   467
  by (subst interval_integral_Icc, simp) (rule set_borel_integral_eq_integral)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   468
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   469
lemma interval_integral_eq_integral':
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   470
  fixes f :: "real \<Rightarrow> 'a::euclidean_space"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   471
  shows "a \<le> b \<Longrightarrow> set_integrable lborel (einterval a b) f \<Longrightarrow> LBINT x=a..b. f x = integral (einterval a b) f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   472
  by (subst interval_lebesgue_integral_le_eq, simp) (rule set_borel_integral_eq_integral)
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   473
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   474
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   475
subsection\<open>General limit approximation arguments\<close>
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   476
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   477
proposition interval_integral_Icc_approx_nonneg:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   478
  fixes a b :: ereal
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   479
  assumes "a < b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   480
  fixes u l :: "nat \<Rightarrow> real"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   481
  assumes  approx: "einterval a b = (\<Union>i. {l i .. u i})"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   482
    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   483
    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   484
  fixes f :: "real \<Rightarrow> real"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   485
  assumes f_integrable: "\<And>i. set_integrable lborel {l i..u i} f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   486
  assumes f_nonneg: "AE x in lborel. a < ereal x \<longrightarrow> ereal x < b \<longrightarrow> 0 \<le> f x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   487
  assumes f_measurable: "set_borel_measurable lborel (einterval a b) f"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   488
  assumes lbint_lim: "(\<lambda>i. LBINT x=l i.. u i. f x) \<longlonglongrightarrow> C"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   489
  shows
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   490
    "set_integrable lborel (einterval a b) f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   491
    "(LBINT x=a..b. f x) = C"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   492
proof -
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   493
  have 1 [unfolded set_integrable_def]: "\<And>i. set_integrable lborel {l i..u i} f" by (rule f_integrable)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   494
  have 2: "AE x in lborel. mono (\<lambda>n. indicator {l n..u 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
   495
  proof -
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   496
     from f_nonneg have "AE x in lborel. \<forall>i. l i \<le> x \<longrightarrow> x \<le> u 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
   497
      by eventually_elim
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   498
         (metis approx(5) approx(6) dual_order.strict_trans1 ereal_less_eq(3) le_less_trans)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   499
    then show ?thesis
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   500
      apply eventually_elim
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   501
      apply (auto simp: mono_def split: split_indicator)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   502
      apply (metis approx(3) decseqD order_trans)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   503
      apply (metis approx(2) incseqD order_trans)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   504
      done
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   505
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   506
  have 3: "AE x in lborel. (\<lambda>i. indicator {l i..u i} x *\<^sub>R f x) \<longlonglongrightarrow> indicator (einterval a b) x *\<^sub>R f x"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   507
  proof -
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   508
    { fix x i assume "l i \<le> x" "x \<le> u i"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   509
      then have "eventually (\<lambda>i. l i \<le> x \<and> x \<le> u i) sequentially"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   510
        apply (auto simp: eventually_sequentially intro!: exI[of _ i])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   511
        apply (metis approx(3) decseqD order_trans)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   512
        apply (metis approx(2) incseqD order_trans)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   513
        done
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   514
      then have "eventually (\<lambda>i. f x * indicator {l i..u i} x = f x) sequentially"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   515
        by eventually_elim auto }
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   516
    then show ?thesis
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   517
      unfolding approx(1) by (auto intro!: AE_I2 tendsto_eventually split: split_indicator)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   518
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   519
  have 4: "(\<lambda>i. \<integral> x. indicator {l i..u i} x *\<^sub>R f x \<partial>lborel) \<longlonglongrightarrow> C"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   520
    using lbint_lim by (simp add: interval_integral_Icc [unfolded set_lebesgue_integral_def] approx less_imp_le)
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   521
  have 5: "(\<lambda>x. indicat_real (einterval a b) x *\<^sub>R f x) \<in> borel_measurable lborel"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   522
    using f_measurable set_borel_measurable_def by blast
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   523
  have "(LBINT x=a..b. f x) = lebesgue_integral lborel (\<lambda>x. indicator (einterval a b) x *\<^sub>R f x)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   524
    using assms by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def less_imp_le)
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   525
  also have "\<dots> = C"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   526
    by (rule integral_monotone_convergence [OF 1 2 3 4 5])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   527
  finally show "(LBINT x=a..b. f x) = C" .
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   528
  show "set_integrable lborel (einterval a b) f"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   529
    unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   530
    by (rule integrable_monotone_convergence[OF 1 2 3 4 5])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   531
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   532
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   533
proposition interval_integral_Icc_approx_integrable:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   534
  fixes u l :: "nat \<Rightarrow> real" and a b :: ereal
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   535
  fixes f :: "real \<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
   536
  assumes "a < b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   537
  assumes  approx: "einterval a b = (\<Union>i. {l i .. u i})"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   538
    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   539
    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   540
  assumes f_integrable: "set_integrable lborel (einterval a b) f"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   541
  shows "(\<lambda>i. LBINT x=l i.. u i. f x) \<longlonglongrightarrow> (LBINT x=a..b. f x)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   542
proof -
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   543
  have "(\<lambda>i. LBINT x:{l i.. u i}. f x) \<longlonglongrightarrow> (LBINT x:einterval a b. f x)"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   544
    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
   545
  proof (rule integral_dominated_convergence)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   546
    show "integrable lborel (\<lambda>x. norm (indicator (einterval a b) x *\<^sub>R f x))"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   547
      using f_integrable integrable_norm set_integrable_def by blast
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   548
    show "(\<lambda>x. indicat_real (einterval a b) x *\<^sub>R f x) \<in> borel_measurable lborel"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   549
      using f_integrable by (simp add: set_integrable_def)
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   550
    then show "\<And>i. (\<lambda>x. indicat_real {l i..u i} x *\<^sub>R f x) \<in> borel_measurable lborel"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   551
      by (rule set_borel_measurable_subset [unfolded set_borel_measurable_def]) (auto simp: approx)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   552
    show "\<And>i. AE x in lborel. norm (indicator {l i..u i} x *\<^sub>R f x) \<le> norm (indicator (einterval a 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
   553
      by (intro AE_I2) (auto simp: approx split: split_indicator)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   554
    show "AE x in lborel. (\<lambda>i. indicator {l i..u i} x *\<^sub>R f x) \<longlonglongrightarrow> indicator (einterval a b) x *\<^sub>R f x"
70365
4df0628e8545 a few new lemmas and a bit of tidying
paulson <lp15@cam.ac.uk>
parents: 70136
diff changeset
   555
    proof (intro AE_I2 tendsto_intros tendsto_eventually)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   556
      fix x
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   557
      { fix i assume "l i \<le> x" "x \<le> u i"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   558
        with \<open>incseq u\<close>[THEN incseqD, of i] \<open>decseq l\<close>[THEN decseqD, of i]
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   559
        have "eventually (\<lambda>i. l i \<le> x \<and> x \<le> u i) sequentially"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   560
          by (auto simp: eventually_sequentially decseq_def incseq_def intro: order_trans) }
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   561
      then show "eventually (\<lambda>xa. indicator {l xa..u xa} x = (indicator (einterval a b) x::real)) sequentially"
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   562
        using approx order_tendstoD(2)[OF \<open>l \<longlonglongrightarrow> a\<close>, of x] order_tendstoD(1)[OF \<open>u \<longlonglongrightarrow> b\<close>, of x]
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   563
        by (auto split: split_indicator)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   564
    qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   565
  qed
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   566
  with \<open>a < b\<close> \<open>\<And>i. l i < u i\<close> show ?thesis
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   567
    by (simp add: interval_lebesgue_integral_le_eq[symmetric] interval_integral_Icc less_imp_le)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   568
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   569
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   570
subsection\<open>A slightly stronger Fundamental Theorem of Calculus\<close>
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   571
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   572
text\<open>Three versions: first, for finite intervals, and then two versions for
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   573
    arbitrary intervals.\<close>
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   574
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   575
(*
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   576
  TODO: make the older versions corollaries of these (using continuous_at_imp_continuous_on, etc.)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   577
*)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   578
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   579
lemma interval_integral_FTC_finite:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   580
  fixes f F :: "real \<Rightarrow> 'a::euclidean_space" and a b :: real
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   581
  assumes f: "continuous_on {min a b..max a b} f"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   582
  assumes F: "\<And>x. min a b \<le> x \<Longrightarrow> x \<le> max a b \<Longrightarrow> (F has_vector_derivative (f x)) (at x within
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   583
    {min a b..max a b})"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   584
  shows "(LBINT x=a..b. f x) = F b - F a"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   585
proof (cases "a \<le> b")
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   586
  case True
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   587
  have "(LBINT x=a..b. f x) = (LBINT x. indicat_real {a..b} x *\<^sub>R f x)"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   588
    by (simp add: True interval_integral_Icc set_lebesgue_integral_def)
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   589
  also have "\<dots> = F b - F a"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   590
  proof (rule integral_FTC_atLeastAtMost [OF True])
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   591
    show "continuous_on {a..b} f"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   592
      using True f by linarith
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   593
    show "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> (F has_vector_derivative f x) (at x within {a..b})"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   594
      by (metis F True max.commute max_absorb1 min_def)
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   595
  qed
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   596
  finally show ?thesis .
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   597
next
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   598
  case False
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   599
  then have "b \<le> a"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   600
    by simp
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   601
  have "- interval_lebesgue_integral lborel (ereal b) (ereal a) f = - (LBINT x. indicat_real {b..a} x *\<^sub>R f x)"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   602
    by (simp add: \<open>b \<le> a\<close> interval_integral_Icc set_lebesgue_integral_def)
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   603
  also have "\<dots> = F b - F a"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   604
  proof (subst integral_FTC_atLeastAtMost [OF \<open>b \<le> a\<close>])
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   605
    show "continuous_on {b..a} f"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   606
      using False f by linarith
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   607
    show "\<And>x. \<lbrakk>b \<le> x; x \<le> a\<rbrakk>
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   608
         \<Longrightarrow> (F has_vector_derivative f x) (at x within {b..a})"
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   609
      by (metis F False max_def min_def)
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   610
  qed auto
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   611
  finally show ?thesis
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   612
    by (metis interval_integral_endpoints_reverse)
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   613
qed
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   614
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   615
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   616
lemma interval_integral_FTC_nonneg:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   617
  fixes f F :: "real \<Rightarrow> real" and a b :: ereal
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   618
  assumes "a < b"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   619
  assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV F x :> f x"
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   620
  assumes f: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f x"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   621
  assumes f_nonneg: "AE x in lborel. a < ereal x \<longrightarrow> ereal x < b \<longrightarrow> 0 \<le> f x"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   622
  assumes A: "((F \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   623
  assumes B: "((F \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   624
  shows
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   625
    "set_integrable lborel (einterval a b) f"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   626
    "(LBINT x=a..b. f x) = B - A"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   627
proof -
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   628
  obtain u l where approx:
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   629
    "einterval a b = (\<Union>i. {l i .. u i})"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   630
    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   631
    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b" 
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   632
    by (blast intro: einterval_Icc_approximation[OF \<open>a < b\<close>])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   633
  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   634
    by (rule order_less_le_trans, rule approx, force)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   635
  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   636
    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   637
  have FTCi: "\<And>i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   638
    using assms approx apply (intro interval_integral_FTC_finite)
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   639
    apply (auto simp: less_imp_le min_def max_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   640
      has_field_derivative_iff_has_vector_derivative[symmetric])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   641
    apply (rule continuous_at_imp_continuous_on, auto intro!: f)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   642
    by (rule DERIV_subset [OF F], auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   643
  have 1: "\<And>i. set_integrable lborel {l i..u i} f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   644
  proof -
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   645
    fix i show "set_integrable lborel {l i .. u i} f"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   646
      using \<open>a < l i\<close> \<open>u i < b\<close> unfolding set_integrable_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   647
      by (intro borel_integrable_compact f continuous_at_imp_continuous_on compact_Icc ballI)
68403
223172b97d0b reorient -> split; documented split
nipkow
parents: 68096
diff changeset
   648
         (auto simp flip: ereal_less_eq)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   649
  qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   650
  have 2: "set_borel_measurable lborel (einterval a b) f"
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   651
    unfolding set_borel_measurable_def
66164
2d79288b042c New theorems and much tidying up of the old ones
paulson <lp15@cam.ac.uk>
parents: 63941
diff changeset
   652
    by (auto simp del: real_scaleR_def intro!: borel_measurable_continuous_on_indicator
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   653
             simp: continuous_on_eq_continuous_at einterval_iff f)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   654
  have 3: "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> B - A"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   655
    apply (subst FTCi)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   656
    apply (intro tendsto_intros)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   657
    using B approx unfolding tendsto_at_iff_sequentially comp_def
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   658
    using tendsto_at_iff_sequentially[where 'a=real]
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   659
    apply (elim allE[of _ "\<lambda>i. ereal (u i)"], auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   660
    using A approx unfolding tendsto_at_iff_sequentially comp_def
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   661
    by (elim allE[of _ "\<lambda>i. ereal (l i)"], auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   662
  show "(LBINT x=a..b. f x) = B - A"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   663
    by (rule interval_integral_Icc_approx_nonneg [OF \<open>a < b\<close> approx 1 f_nonneg 2 3])
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   664
  show "set_integrable lborel (einterval a b) f"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   665
    by (rule interval_integral_Icc_approx_nonneg [OF \<open>a < b\<close> approx 1 f_nonneg 2 3])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   666
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   667
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   668
theorem interval_integral_FTC_integrable:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   669
  fixes f F :: "real \<Rightarrow> 'a::euclidean_space" and a b :: ereal
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   670
  assumes "a < b"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   671
  assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> (F has_vector_derivative f x) (at x)"
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   672
  assumes f: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f x"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   673
  assumes f_integrable: "set_integrable lborel (einterval a b) f"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   674
  assumes A: "((F \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   675
  assumes B: "((F \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   676
  shows "(LBINT x=a..b. f x) = B - A"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   677
proof -
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   678
  obtain u l where approx:
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   679
    "einterval a b = (\<Union>i. {l i .. u i})"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   680
    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   681
    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b" 
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   682
    by (blast intro: einterval_Icc_approximation[OF \<open>a < b\<close>])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   683
  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   684
    by (rule order_less_le_trans, rule approx, force)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   685
  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   686
    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   687
  have FTCi: "\<And>i. (LBINT x=l i..u i. f x) = F (u i) - F (l i)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   688
    using assms approx
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   689
    by (auto simp: less_imp_le min_def max_def
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   690
             intro!: f continuous_at_imp_continuous_on interval_integral_FTC_finite
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   691
             intro: has_vector_derivative_at_within)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   692
  have "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> B - A"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   693
    unfolding FTCi
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   694
  proof (intro tendsto_intros)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   695
    show "(\<lambda>x. F (l x)) \<longlonglongrightarrow> A"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   696
      using A approx unfolding tendsto_at_iff_sequentially comp_def
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   697
      by (elim allE[of _ "\<lambda>i. ereal (l i)"], auto)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   698
    show "(\<lambda>x. F (u x)) \<longlonglongrightarrow> B"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   699
      using B approx unfolding tendsto_at_iff_sequentially comp_def
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   700
      by (elim allE[of _ "\<lambda>i. ereal (u i)"], auto)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   701
  qed
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   702
  moreover have "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> (LBINT x=a..b. f x)"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   703
    by (rule interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx f_integrable])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   704
  ultimately show ?thesis
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   705
    by (elim LIMSEQ_unique)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   706
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   707
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   708
(*
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   709
  The second Fundamental Theorem of Calculus and existence of antiderivatives on an
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   710
  einterval.
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   711
*)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   712
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   713
theorem interval_integral_FTC2:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   714
  fixes a b c :: real and f :: "real \<Rightarrow> 'a::euclidean_space"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   715
  assumes "a \<le> c" "c \<le> b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   716
  and contf: "continuous_on {a..b} f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   717
  fixes x :: real
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   718
  assumes "a \<le> x" and "x \<le> b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   719
  shows "((\<lambda>u. LBINT y=c..u. f y) has_vector_derivative (f x)) (at x within {a..b})"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   720
proof -
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   721
  let ?F = "(\<lambda>u. LBINT y=a..u. f y)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   722
  have intf: "set_integrable lborel {a..b} f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   723
    by (rule borel_integrable_atLeastAtMost', rule contf)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   724
  have "((\<lambda>u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   725
    using \<open>a \<le> x\<close> \<open>x \<le> b\<close> 
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   726
    by (auto intro: integral_has_vector_derivative continuous_on_subset [OF contf])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   727
  then have "((\<lambda>u. integral {a..u} f) has_vector_derivative (f x)) (at x within {a..b})"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   728
    by simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   729
  then have "(?F has_vector_derivative (f x)) (at x within {a..b})"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   730
    by (rule has_vector_derivative_weaken)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   731
       (auto intro!: assms interval_integral_eq_integral[symmetric] set_integrable_subset [OF intf])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   732
  then have "((\<lambda>x. (LBINT y=c..a. f y) + ?F x) has_vector_derivative (f x)) (at x within {a..b})"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   733
    by (auto intro!: derivative_eq_intros)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   734
  then show ?thesis
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   735
  proof (rule has_vector_derivative_weaken)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   736
    fix u assume "u \<in> {a .. b}"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   737
    then show "(LBINT y=c..a. f y) + (LBINT y=a..u. f y) = (LBINT y=c..u. f y)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   738
      using assms
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   739
      apply (intro interval_integral_sum)
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   740
      apply (auto simp: interval_lebesgue_integrable_def simp del: real_scaleR_def)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   741
      by (rule set_integrable_subset [OF intf], auto simp: min_def max_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   742
  qed (insert assms, auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   743
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   744
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   745
proposition einterval_antiderivative:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   746
  fixes a b :: ereal and f :: "real \<Rightarrow> 'a::euclidean_space"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   747
  assumes "a < b" and contf: "\<And>x :: real. a < x \<Longrightarrow> x < b \<Longrightarrow> isCont f x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   748
  shows "\<exists>F. \<forall>x :: real. a < x \<longrightarrow> x < b \<longrightarrow> (F has_vector_derivative f x) (at x)"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   749
proof -
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   750
  from einterval_nonempty [OF \<open>a < b\<close>] obtain c :: real where [simp]: "a < c" "c < b"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   751
    by (auto simp: einterval_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   752
  let ?F = "(\<lambda>u. LBINT y=c..u. f y)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   753
  show ?thesis
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   754
  proof (rule exI, clarsimp)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   755
    fix x :: real
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   756
    assume [simp]: "a < x" "x < b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   757
    have 1: "a < min c x" by simp
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   758
    from einterval_nonempty [OF 1] obtain d :: real where [simp]: "a < d" "d < c" "d < x"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   759
      by (auto simp: einterval_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   760
    have 2: "max c x < b" by simp
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   761
    from einterval_nonempty [OF 2] obtain e :: real where [simp]: "c < e" "x < e" "e < b"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   762
      by (auto simp: einterval_def)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   763
    have "(?F has_vector_derivative f x) (at x within {d<..<e})"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   764
    proof (rule has_vector_derivative_within_subset [of _ _ _ "{d..e}"])
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   765
      have "continuous_on {d..e} f"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   766
      proof (intro continuous_at_imp_continuous_on ballI contf; clarsimp)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   767
        show "\<And>x. \<lbrakk>d \<le> x; x \<le> e\<rbrakk> \<Longrightarrow> a < ereal x"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   768
          using \<open>a < ereal d\<close> ereal_less_ereal_Ex by auto
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   769
        show "\<And>x. \<lbrakk>d \<le> x; x \<le> e\<rbrakk> \<Longrightarrow> ereal x < b"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   770
          using \<open>ereal e < b\<close> ereal_less_eq(3) le_less_trans by blast
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   771
      qed
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   772
      then show "(?F has_vector_derivative f x) (at x within {d..e})"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   773
        by (intro interval_integral_FTC2) (use \<open>d < c\<close> \<open>c < e\<close> \<open>d < x\<close> \<open>x < e\<close> in \<open>linarith+\<close>)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   774
    qed auto
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   775
    then show "(?F has_vector_derivative f x) (at x)"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   776
      by (force simp: has_vector_derivative_within_open [of _ "{d<..<e}"])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   777
  qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   778
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   779
69683
8b3458ca0762 subsection is always %important
immler
parents: 69681
diff changeset
   780
subsection\<open>The substitution theorem\<close>
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   781
67974
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   782
text\<open>Once again, three versions: first, for finite intervals, and then two versions for
3f352a91b45a replacement of set integral abbreviations by actual definitions!
paulson <lp15@cam.ac.uk>
parents: 66164
diff changeset
   783
    arbitrary intervals.\<close>
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   784
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   785
theorem interval_integral_substitution_finite:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   786
  fixes a b :: real and f :: "real \<Rightarrow> 'a::euclidean_space"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   787
  assumes "a \<le> b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   788
  and derivg: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (g has_real_derivative (g' x)) (at x within {a..b})"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   789
  and contf : "continuous_on (g ` {a..b}) f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   790
  and contg': "continuous_on {a..b} g'"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   791
  shows "LBINT x=a..b. g' x *\<^sub>R f (g x) = LBINT y=g a..g b. f y"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   792
proof-
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   793
  have v_derivg: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (g has_vector_derivative (g' x)) (at x within {a..b})"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   794
    using derivg unfolding has_field_derivative_iff_has_vector_derivative .
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   795
  then have contg [simp]: "continuous_on {a..b} g"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   796
    by (rule continuous_on_vector_derivative) auto
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   797
  have 1: "\<exists>x\<in>{a..b}. u = g x" if "min (g a) (g b) \<le> u" "u \<le> max (g a) (g b)" for u
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   798
    by (cases "g a \<le> g b") (use that assms IVT' [of g a u b]  IVT2' [of g b u a]  in \<open>auto simp: min_def max_def\<close>)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   799
  obtain c d where g_im: "g ` {a..b} = {c..d}" and "c \<le> d"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   800
    by (metis continuous_image_closed_interval contg \<open>a \<le> b\<close>)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   801
  obtain F where derivF:
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   802
         "\<And>x. \<lbrakk>a \<le> x; x \<le> b\<rbrakk> \<Longrightarrow> (F has_vector_derivative (f (g x))) (at (g x) within (g ` {a..b}))" 
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   803
    using continuous_on_subset [OF contf] g_im
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   804
      by (metis antiderivative_continuous atLeastAtMost_iff image_subset_iff set_eq_subset)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   805
  have contfg: "continuous_on {a..b} (\<lambda>x. f (g x))"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   806
    by (blast intro: continuous_on_compose2 contf contg)
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   807
  have "LBINT x. indicat_real {a..b} x *\<^sub>R g' x *\<^sub>R f (g x) = F (g b) - F (g a)"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   808
    apply (rule integral_FTC_atLeastAtMost
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   809
                [OF \<open>a \<le> b\<close> vector_diff_chain_within[OF v_derivg derivF, unfolded comp_def]])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   810
    apply (auto intro!: continuous_on_scaleR contg' contfg)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   811
    done
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   812
  then have "LBINT x=a..b. g' x *\<^sub>R f (g x) = F (g b) - F (g a)"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   813
    by (simp add: assms interval_integral_Icc set_lebesgue_integral_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   814
  moreover have "LBINT y=(g a)..(g b). f y = F (g b) - F (g a)"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   815
  proof (rule interval_integral_FTC_finite)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   816
    show "continuous_on {min (g a) (g b)..max (g a) (g b)} f"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   817
      by (rule continuous_on_subset [OF contf]) (auto simp: image_def 1)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   818
    show "(F has_vector_derivative f y) (at y within {min (g a) (g b)..max (g a) (g b)})" 
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   819
      if y: "min (g a) (g b) \<le> y" "y \<le> max (g a) (g b)" for y
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   820
    proof -
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   821
      obtain x where "a \<le> x" "x \<le> b" "y = g x"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   822
        using 1 y by force
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   823
      then show ?thesis
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   824
        by (auto simp: image_def intro!: 1  has_vector_derivative_within_subset [OF derivF])
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   825
    qed
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   826
  qed
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   827
  ultimately show ?thesis by simp
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   828
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   829
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   830
(* TODO: is it possible to lift the assumption here that g' is nonnegative? *)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   831
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   832
theorem interval_integral_substitution_integrable:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   833
  fixes f :: "real \<Rightarrow> 'a::euclidean_space" and a b u v :: ereal
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   834
  assumes "a < b"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   835
  and deriv_g: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV g x :> g' x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   836
  and contf: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f (g x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   837
  and contg': "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont g' x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   838
  and g'_nonneg: "\<And>x. a \<le> ereal x \<Longrightarrow> ereal x \<le> b \<Longrightarrow> 0 \<le> g' x"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   839
  and A: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   840
  and B: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   841
  and integrable: "set_integrable lborel (einterval a b) (\<lambda>x. g' x *\<^sub>R f (g x))"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   842
  and integrable2: "set_integrable lborel (einterval A B) (\<lambda>x. f x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   843
  shows "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   844
proof -
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   845
  obtain u l where approx [simp]:
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   846
    "einterval a b = (\<Union>i. {l i .. u i})"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   847
    "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   848
    "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b" 
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   849
    by (blast intro: einterval_Icc_approximation[OF \<open>a < b\<close>])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   850
  note less_imp_le [simp]
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   851
  have [simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   852
    by (rule order_less_le_trans, rule approx, force)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   853
  have [simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   854
    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   855
  then have lessb[simp]: "\<And>i. l i < b"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   856
    using approx(4) less_eq_real_def by blast
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   857
  have [simp]: "\<And>i. a < u i"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   858
    by (rule order_less_trans, rule approx, auto, rule approx)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   859
  have lle[simp]: "\<And>i j. i \<le> j \<Longrightarrow> l j \<le> l i" by (rule decseqD, rule approx)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   860
  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> u i \<le> u j" by (rule incseqD, rule approx)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   861
  have g_nondec [simp]: "g x \<le> g y" if "a < x" "x \<le> y" "y < b" for x y
68638
87d1bff264df de-applying and meta-quantifying
paulson <lp15@cam.ac.uk>
parents: 68532
diff changeset
   862
  proof (rule DERIV_nonneg_imp_nondecreasing [OF \<open>x \<le> y\<close>], intro exI conjI)
87d1bff264df de-applying and meta-quantifying
paulson <lp15@cam.ac.uk>
parents: 68532
diff changeset
   863
    show "\<And>u. x \<le> u \<Longrightarrow> u \<le> y \<Longrightarrow> (g has_real_derivative g' u) (at u)"
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   864
      by (meson deriv_g ereal_less_eq(3) le_less_trans less_le_trans that)
68638
87d1bff264df de-applying and meta-quantifying
paulson <lp15@cam.ac.uk>
parents: 68532
diff changeset
   865
    show "\<And>u. x \<le> u \<Longrightarrow> u \<le> y \<Longrightarrow> 0 \<le> g' u"
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   866
      by (meson assms(5) dual_order.trans le_ereal_le less_imp_le order_refl that)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   867
  qed
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   868
  have "A \<le> B" and un: "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   869
  proof -
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   870
    have A2: "(\<lambda>i. g (l i)) \<longlonglongrightarrow> A"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   871
      using A apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   872
      by (drule_tac x = "\<lambda>i. ereal (l i)" in spec, auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   873
    hence A3: "\<And>i. g (l i) \<ge> A"
68532
f8b98d31ad45 Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents: 68403
diff changeset
   874
      by (intro decseq_ge, auto simp: decseq_def)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   875
    have B2: "(\<lambda>i. g (u i)) \<longlonglongrightarrow> B"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   876
      using B apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   877
      by (drule_tac x = "\<lambda>i. ereal (u i)" in spec, auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   878
    hence B3: "\<And>i. g (u i) \<le> B"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   879
      by (intro incseq_le, auto simp: incseq_def)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   880
    have "ereal (g (l 0)) \<le> ereal (g (u 0))"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   881
      by auto
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   882
    then show "A \<le> B"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   883
      by (meson A3 B3 order.trans)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   884
    { fix x :: real
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   885
      assume "A < x" and "x < B"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   886
      then have "eventually (\<lambda>i. ereal (g (l i)) < x \<and> x < ereal (g (u i))) sequentially"
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   887
        by (fast intro: eventually_conj order_tendstoD A2 B2)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   888
      hence "\<exists>i. g (l i) < x \<and> x < g (u i)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   889
        by (simp add: eventually_sequentially, auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   890
    } note AB = this
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   891
    show "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   892
    proof
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   893
      show "einterval A B \<subseteq> (\<Union>i. {g(l i)<..<g(u i)})"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   894
        by (auto simp: einterval_def AB)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   895
      show "(\<Union>i. {g(l i)<..<g(u i)}) \<subseteq> einterval A B"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   896
      proof (clarsimp simp add: einterval_def, intro conjI)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   897
        show "\<And>x i. \<lbrakk>g (l i) < x; x < g (u i)\<rbrakk> \<Longrightarrow> A < ereal x"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   898
          using A3 le_ereal_less by blast
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   899
        show "\<And>x i. \<lbrakk>g (l i) < x; x < g (u i)\<rbrakk> \<Longrightarrow> ereal x < B"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   900
          using B3 ereal_le_less by blast
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   901
      qed
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   902
    qed
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   903
  qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   904
  (* finally, the main argument *)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   905
  have eq1: "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)" for i
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   906
    apply (rule interval_integral_substitution_finite [OF _ DERIV_subset [OF deriv_g]])
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   907
    unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   908
         apply (auto intro!: continuous_at_imp_continuous_on contf contg')
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   909
    done
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   910
  have "(\<lambda>i. LBINT x=l i..u i. g' x *\<^sub>R f (g x)) \<longlonglongrightarrow> (LBINT x=a..b. g' x *\<^sub>R f (g x))"
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   911
    apply (rule interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx])
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   912
    by (rule assms)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   913
  hence 2: "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> (LBINT x=a..b. g' x *\<^sub>R f (g x))"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   914
    by (simp add: eq1)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   915
  have incseq: "incseq (\<lambda>i. {g (l i)<..<g (u i)})"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   916
    apply (auto simp: incseq_def)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   917
    using lessb lle approx(5) g_nondec le_less_trans apply blast
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   918
    by (force intro: less_le_trans)
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   919
  have "(\<lambda>i. set_lebesgue_integral lborel {g (l i)<..<g (u i)} f)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   920
        \<longlonglongrightarrow> set_lebesgue_integral lborel (einterval A B) f"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   921
    unfolding un  by (rule set_integral_cont_up) (use incseq  integrable2 un in auto)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   922
  then have "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> (LBINT x = A..B. f x)"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   923
    by (simp add: interval_lebesgue_integral_le_eq \<open>A \<le> B\<close>)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   924
  thus ?thesis by (intro LIMSEQ_unique [OF _ 2])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   925
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   926
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   927
(* TODO: the last two proofs are only slightly different. Factor out common part?
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   928
   An alternative: make the second one the main one, and then have another lemma
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   929
   that says that if f is nonnegative and all the other hypotheses hold, then it is integrable. *)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   930
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   931
theorem interval_integral_substitution_nonneg:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   932
  fixes f g g':: "real \<Rightarrow> real" and a b u v :: ereal
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   933
  assumes "a < b"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   934
  and deriv_g: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV g x :> g' x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   935
  and contf: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont f (g x)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   936
  and contg': "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> isCont g' x"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   937
  and f_nonneg: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> 0 \<le> f (g x)" (* TODO: make this AE? *)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   938
  and g'_nonneg: "\<And>x. a \<le> ereal x \<Longrightarrow> ereal x \<le> b \<Longrightarrow> 0 \<le> g' x"
61973
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   939
  and A: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)"
0c7e865fa7cb more symbols;
wenzelm
parents: 61969
diff changeset
   940
  and B: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> B) (at_left b)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   941
  and integrable_fg: "set_integrable lborel (einterval a b) (\<lambda>x. f (g x) * g' x)"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   942
  shows
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   943
    "set_integrable lborel (einterval A B) f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   944
    "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
   945
proof -
61808
fc1556774cfe isabelle update_cartouches -c -t;
wenzelm
parents: 61609
diff changeset
   946
  from einterval_Icc_approximation[OF \<open>a < b\<close>] guess u l . note approx [simp] = this
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   947
  note less_imp_le [simp]
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   948
  have aless[simp]: "\<And>x i. l i \<le> x \<Longrightarrow> a < ereal x"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   949
    by (rule order_less_le_trans, rule approx, force)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   950
  have lessb[simp]: "\<And>x i. x \<le> u i \<Longrightarrow> ereal x < b"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   951
    by (rule order_le_less_trans, subst ereal_less_eq(3), assumption, rule approx)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   952
  have llb[simp]: "\<And>i. l i < b"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   953
    using lessb approx(4) less_eq_real_def by blast
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   954
  have alu[simp]: "\<And>i. a < u i"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   955
    by (rule order_less_trans, rule approx, auto, rule approx)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   956
  have [simp]: "\<And>i j. i \<le> j \<Longrightarrow> l j \<le> l i" by (rule decseqD, rule approx)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   957
  have uleu[simp]: "\<And>i j. i \<le> j \<Longrightarrow> u i \<le> u j" by (rule incseqD, rule approx)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   958
  have g_nondec [simp]: "g x \<le> g y" if "a < x" "x \<le> y" "y < b" for x y
68638
87d1bff264df de-applying and meta-quantifying
paulson <lp15@cam.ac.uk>
parents: 68532
diff changeset
   959
  proof (rule DERIV_nonneg_imp_nondecreasing [OF \<open>x \<le> y\<close>], intro exI conjI)
87d1bff264df de-applying and meta-quantifying
paulson <lp15@cam.ac.uk>
parents: 68532
diff changeset
   960
    show "\<And>u. x \<le> u \<Longrightarrow> u \<le> y \<Longrightarrow> (g has_real_derivative g' u) (at u)"
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   961
      by (meson deriv_g ereal_less_eq(3) le_less_trans less_le_trans that)
68638
87d1bff264df de-applying and meta-quantifying
paulson <lp15@cam.ac.uk>
parents: 68532
diff changeset
   962
    show "\<And>u. x \<le> u \<Longrightarrow> u \<le> y \<Longrightarrow> 0 \<le> g' u"
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   963
      by (meson g'_nonneg less_ereal.simps(1) less_trans not_less that)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   964
  qed
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   965
  have "A \<le> B" and un: "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   966
  proof -
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   967
    have A2: "(\<lambda>i. g (l i)) \<longlonglongrightarrow> A"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   968
      using A apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   969
      by (drule_tac x = "\<lambda>i. ereal (l i)" in spec, auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   970
    hence A3: "\<And>i. g (l i) \<ge> A"
68532
f8b98d31ad45 Incorporating new/strengthened proofs from Library and AFP entries
paulson <lp15@cam.ac.uk>
parents: 68403
diff changeset
   971
      by (intro decseq_ge, auto simp: decseq_def)
61969
e01015e49041 more symbols;
wenzelm
parents: 61945
diff changeset
   972
    have B2: "(\<lambda>i. g (u i)) \<longlonglongrightarrow> B"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   973
      using B apply (auto simp: einterval_def tendsto_at_iff_sequentially comp_def)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   974
      by (drule_tac x = "\<lambda>i. ereal (u i)" in spec, auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   975
    hence B3: "\<And>i. g (u i) \<le> B"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   976
      by (intro incseq_le, auto simp: incseq_def)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   977
    have "ereal (g (l 0)) \<le> ereal (g (u 0))"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   978
      by auto
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   979
    then show "A \<le> B"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   980
      by (meson A3 B3 order.trans)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   981
    { fix x :: real
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
   982
      assume "A < x" and "x < B"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   983
      then have "eventually (\<lambda>i. ereal (g (l i)) < x \<and> x < ereal (g (u i))) sequentially"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   984
        by (fast intro: eventually_conj order_tendstoD A2 B2)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   985
      hence "\<exists>i. g (l i) < x \<and> x < g (u i)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   986
        by (simp add: eventually_sequentially, auto)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   987
    } note AB = this
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   988
    show "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   989
    proof
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   990
      show "einterval A B \<subseteq> (\<Union>i. {g (l i)<..<g (u i)})"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
   991
        by (auto simp: einterval_def AB)
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   992
      show "(\<Union>i. {g (l i)<..<g (u i)}) \<subseteq> einterval A B"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   993
        apply (clarsimp simp: einterval_def, intro conjI)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   994
        using A3 le_ereal_less apply blast
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   995
        using B3 ereal_le_less by blast
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   996
    qed
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
   997
  qed
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   998
    (* finally, the main argument *)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
   999
  have eq1: "(LBINT x=l i.. u i. (f (g x) * g' x)) = (LBINT y=g (l i)..g (u i). f y)" for i
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1000
  proof -
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1001
    have "(LBINT x=l i.. u i. g' x *\<^sub>R f (g x)) = (LBINT y=g (l i)..g (u i). f y)"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1002
      apply (rule interval_integral_substitution_finite [OF _ DERIV_subset [OF deriv_g]])
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1003
      unfolding has_field_derivative_iff_has_vector_derivative[symmetric]
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1004
           apply (auto intro!: continuous_at_imp_continuous_on contf contg')
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1005
      done
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1006
    then show ?thesis
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1007
      by (simp add: ac_simps)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1008
  qed
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1009
  have incseq: "incseq (\<lambda>i. {g (l i)<..<g (u i)})"
68095
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1010
    apply (clarsimp simp add: incseq_def, intro conjI)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1011
    apply (meson llb antimono_def approx(3) approx(5) g_nondec le_less_trans)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1012
    using alu uleu approx(6) g_nondec less_le_trans by blast
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1013
  have img: "\<exists>c \<ge> l i. c \<le> u i \<and> x = g c" if "g (l i) \<le> x" "x \<le> g (u i)" for x i
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1014
  proof -
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1015
    have "continuous_on {l i..u i} g"
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1016
      by (force intro!: DERIV_isCont deriv_g continuous_at_imp_continuous_on)
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1017
    with that show ?thesis
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1018
      using IVT' [of g] approx(4) dual_order.strict_implies_order by blast
4fa3e63ecc7e starting to tidy up Interval_Integral.thy
paulson <lp15@cam.ac.uk>
parents: 68046
diff changeset
  1019
  qed
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1020
  have "continuous_on {g (l i)..g (u i)} f" for i
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1021
    apply (intro continuous_intros continuous_at_imp_continuous_on)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1022
    using contf img by force
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1023
  then have int_f: "\<And>i. set_integrable lborel {g (l i)<..<g (u i)} f"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1024
    by (rule set_integrable_subset [OF borel_integrable_atLeastAtMost']) (auto intro: less_imp_le)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1025
  have integrable: "set_integrable lborel (\<Union>i. {g (l i)<..<g (u i)}) f"
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1026
  proof (intro pos_integrable_to_top incseq int_f)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1027
    let ?l = "(LBINT x=a..b. f (g x) * g' x)"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1028
    have "(\<lambda>i. LBINT x=l i..u i. f (g x) * g' x) \<longlonglongrightarrow> ?l"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1029
      by (intro assms interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx])
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1030
    hence "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> ?l"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1031
      by (simp add: eq1)
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1032
    then show "(\<lambda>i. set_lebesgue_integral lborel {g (l i)<..<g (u i)} f) \<longlonglongrightarrow> ?l"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1033
      unfolding interval_lebesgue_integral_def by auto
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1034
    have "\<And>x i. g (l i) \<le> x \<Longrightarrow> x \<le> g (u i) \<Longrightarrow> 0 \<le> f x"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1035
      using aless f_nonneg img lessb by blast
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1036
    then show "\<And>x i. x \<in> {g (l i)<..<g (u i)} \<Longrightarrow> 0 \<le> f x"
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1037
      using less_eq_real_def by auto
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1038
  qed (auto simp: greaterThanLessThan_borel)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1039
  thus "set_integrable lborel (einterval A B) f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1040
    by (simp add: un)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1041
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1042
  have "(LBINT x=A..B. f x) = (LBINT x=a..b. g' x *\<^sub>R f (g x))"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1043
  proof (rule interval_integral_substitution_integrable)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1044
    show "set_integrable lborel (einterval a b) (\<lambda>x. g' x *\<^sub>R f (g x))"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1045
      using integrable_fg by (simp add: ac_simps)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1046
  qed fact+
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1047
  then show "(LBINT x=A..B. f x) = (LBINT x=a..b. (f (g x) * g' x))"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1048
    by (simp add: ac_simps)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1049
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1050
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1051
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
  1052
syntax "_complex_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real \<Rightarrow> complex"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
  1053
  ("(2CLBINT _. _)" [0,60] 60)
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
  1054
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
  1055
translations "CLBINT x. f" == "CONST complex_lebesgue_integral CONST lborel (\<lambda>x. f)"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
  1056
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
  1057
syntax "_complex_set_lebesgue_borel_integral" :: "pttrn \<Rightarrow> real set \<Rightarrow> real \<Rightarrow> complex"
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
  1058
  ("(3CLBINT _:_. _)" [0,60,61] 60)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1059
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1060
translations
63941
f353674c2528 move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
hoelzl
parents: 63886
diff changeset
  1061
  "CLBINT x:A. f" == "CONST complex_set_lebesgue_integral CONST lborel A (\<lambda>x. f)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1062
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
  1063
abbreviation complex_interval_lebesgue_integral ::
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1064
    "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> complex" where
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1065
  "complex_interval_lebesgue_integral M a b f \<equiv> interval_lebesgue_integral M a b f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1066
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
  1067
abbreviation complex_interval_lebesgue_integrable ::
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1068
  "real measure \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> (real \<Rightarrow> complex) \<Rightarrow> bool" where
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1069
  "complex_interval_lebesgue_integrable M a b f \<equiv> interval_lebesgue_integrable M a b f"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1070
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1071
syntax
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1072
  "_ascii_complex_interval_lebesgue_borel_integral" :: "pttrn \<Rightarrow> ereal \<Rightarrow> ereal \<Rightarrow> real \<Rightarrow> complex"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1073
  ("(4CLBINT _=_.._. _)" [0,60,60,61] 60)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1074
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1075
translations
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1076
  "CLBINT x=a..b. f" == "CONST complex_interval_lebesgue_integral CONST lborel a b (\<lambda>x. f)"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1077
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1078
proposition interval_integral_norm:
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1079
  fixes f :: "real \<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
  1080
  shows "interval_lebesgue_integrable lborel a b f \<Longrightarrow> a \<le> b \<Longrightarrow>
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1081
    norm (LBINT t=a..b. f t) \<le> LBINT t=a..b. norm (f t)"
70136
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69683
diff changeset
  1082
  using integral_norm_bound[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"]
f03a01a18c6e modernized tags: default scope excludes proof;
wenzelm
parents: 69683
diff changeset
  1083
  by (auto simp: interval_lebesgue_integral_def interval_lebesgue_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
  1084
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1085
proposition interval_integral_norm2:
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
  1086
  "interval_lebesgue_integrable lborel a b f \<Longrightarrow>
61945
1135b8de26c3 more symbols;
wenzelm
parents: 61897
diff changeset
  1087
    norm (LBINT t=a..b. f t) \<le> \<bar>LBINT t=a..b. norm (f t)\<bar>"
69681
689997a8a582 redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
immler
parents: 69680
diff changeset
  1088
proof (induct a b rule: linorder_wlog)
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1089
  case (sym a b) then show ?case
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1090
    by (simp add: interval_integral_endpoints_reverse[of a b] interval_integrable_endpoints_reverse[of a b])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1091
next
63329
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
  1092
  case (le a b)
6b26c378ab35 Probability: tuned headers; cleanup Radon_Nikodym
hoelzl
parents: 63092
diff changeset
  1093
  then have "\<bar>LBINT t=a..b. norm (f t)\<bar> = LBINT t=a..b. norm (f t)"
59092
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1094
    using integrable_norm[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"]
68096
e58c9ac761cb more tidying
paulson <lp15@cam.ac.uk>
parents: 68095
diff changeset
  1095
    by (auto simp: interval_lebesgue_integral_def interval_lebesgue_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
  1096
             intro!: integral_nonneg_AE abs_of_nonneg)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1097
  then show ?case
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1098
    using le by (simp add: interval_integral_norm)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1099
qed
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1100
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1101
(* TODO: should we have a library of facts like these? *)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1102
lemma integral_cos: "t \<noteq> 0 \<Longrightarrow> LBINT x=a..b. cos (t * x) = sin (t * b) / t - sin (t * a) / t"
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1103
  apply (intro interval_integral_FTC_finite continuous_intros)
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1104
  by (auto intro!: derivative_eq_intros simp: has_field_derivative_iff_has_vector_derivative[symmetric])
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1105
d469103c0737 add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
hoelzl
parents:
diff changeset
  1106
end