| author | wenzelm | 
| Sat, 13 Aug 2022 11:22:51 +0200 | |
| changeset 75835 | 5c53e24d3dc2 | 
| parent 75462 | 7448423e5dba | 
| child 78480 | b22f39c54e8c | 
| permissions | -rw-r--r-- | 
| 63627 | 1 | (* Title: HOL/Analysis/Interval_Integral.thy | 
| 63329 | 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: 
69680diff
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: 
63886diff
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: 
69680diff
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 | 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 | 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: 
69680diff
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: 
68046diff
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: 
69680diff
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 | 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 | 49 | moreover have "(\<lambda>x. ereal (real (Suc x))) \<longlonglongrightarrow> \<infinity>" | 
| 71827 | 50 | by (simp add: Lim_PInfty filterlim_sequentially_Suc) (metis le_SucI of_nat_Suc of_nat_mono order_trans real_arch_simple) | 
| 61969 | 51 | moreover have "\<And>r. (\<lambda>x. ereal (r + real (Suc x))) \<longlonglongrightarrow> \<infinity>" | 
| 71827 | 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: 
59867diff
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 | 58 | define d where "d = b' - (if a = -\<infinity> then b' - 1 else real_of_ereal a)" | 
| 61808 | 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 | 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: 
68046diff
changeset | 66 | moreover | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
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: 
68046diff
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: 
68046diff
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: 
68046diff
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: 
68046diff
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 | 74 | (auto simp: real incseq_def intro!: divide_left_mono) | 
| 74362 | 75 | qed (use \<open>a < b\<close> in 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: 
69680diff
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 | 80 | obtains X :: "nat \<Rightarrow> real" where | 
| 61969 | 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: 
69680diff
changeset | 82 | proof - | 
| 61808 | 83 | have "-b < -a" using \<open>a < b\<close> by simp | 
| 74362 | 84 | from ereal_incseq_approx[OF this] obtain X where | 
| 85 | "incseq X" | |
| 86 | "\<And>i. - b < ereal (X i)" | |
| 87 | "\<And>i. ereal (X i) < - a" | |
| 88 | "(\<lambda>x. ereal (X x)) \<longlonglongrightarrow> - a" | |
| 89 | by auto | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 90 | then show thesis | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 91 | apply (intro that[of "\<lambda>i. - X i"]) | 
| 68403 | 92 | 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 | 93 | 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 | 94 | done | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 95 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 96 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 97 | proposition einterval_Icc_approximation: | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 98 | fixes a b :: ereal | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 99 | assumes "a < b" | 
| 63329 | 100 | 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 | 101 |     "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 | 102 | "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b" | 
| 61969 | 103 | "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 104 | proof - | 
| 61808 | 105 | from dense[OF \<open>a < b\<close>] obtain c where "a < c" "c < b" by safe | 
| 74362 | 106 | from ereal_incseq_approx[OF \<open>c < b\<close>] obtain u where u: | 
| 107 | "incseq u" | |
| 108 | "\<And>i. c < ereal (u i)" | |
| 109 | "\<And>i. ereal (u i) < b" | |
| 110 | "(\<lambda>x. ereal (u x)) \<longlonglongrightarrow> b" | |
| 111 | by auto | |
| 112 | from ereal_decseq_approx[OF \<open>a < c\<close>] obtain l where l: | |
| 113 | "decseq l" | |
| 114 | "\<And>i. a < ereal (l i)" | |
| 115 | "\<And>i. ereal (l i) < c" | |
| 116 | "(\<lambda>x. ereal (l x)) \<longlonglongrightarrow> a" | |
| 117 | by auto | |
| 61808 | 118 |   { 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 | 119 |   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 | 120 | proof (auto simp: einterval_iff) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 121 | 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 | 122 | have "eventually (\<lambda>i. ereal (l i) < ereal x) sequentially" | 
| 61808 | 123 | using l(4) \<open>a < ereal x\<close> by (rule order_tendstoD) | 
| 63329 | 124 | moreover | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 125 | have "eventually (\<lambda>i. ereal x < ereal (u i)) sequentially" | 
| 61808 | 126 | 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 | 127 | 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 | 128 | by eventually_elim auto | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 129 | 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 | 130 | 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 | 131 | next | 
| 63329 | 132 | fix x i assume "l i \<le> x" "x \<le> u i" | 
| 61808 | 133 | 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 | 134 | show "a < ereal x" "ereal x < b" | 
| 68403 | 135 | 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 | 136 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 137 | show thesis | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 138 | by (intro that) fact+ | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 139 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 140 | |
| 63329 | 141 | (* 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 | 142 | they are real. *) | 
| 70136 | 143 | 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 | 144 | "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 | 145 | (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 | 146 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 147 | syntax | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 148 | "_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 | 149 |   ("(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 | 150 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 151 | translations | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 152 | "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 | 153 | |
| 70136 | 154 | 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 | 155 | "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 | 156 | (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 | 157 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 158 | syntax | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 159 | "_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 | 160 |   ("(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 | 161 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 162 | translations | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 163 | "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 | 164 | |
| 69683 | 165 | 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 | 166 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 167 | lemma interval_lebesgue_integral_cong: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 168 | "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 | 169 | 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 | 170 | 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 | 171 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 172 | 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 | 173 | "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 | 174 | 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 | 175 | 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 | 176 | 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 | 177 | |
| 62083 | 178 | lemma interval_integrable_mirror: | 
| 179 | shows "interval_lebesgue_integrable lborel a b (\<lambda>x. f (-x)) \<longleftrightarrow> | |
| 180 | interval_lebesgue_integrable lborel (-b) (-a) f" | |
| 181 | proof - | |
| 182 | have *: "indicator (einterval a b) (- x) = (indicator (einterval (-b) (-a)) x :: real)" | |
| 183 | for a b :: ereal and x :: real | |
| 184 | by (cases a b rule: ereal2_cases) (auto simp: einterval_def split: split_indicator) | |
| 185 | show ?thesis | |
| 186 | unfolding interval_lebesgue_integrable_def | |
| 187 | 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: 
66164diff
changeset | 188 | by (simp add: * set_integrable_def) | 
| 62083 | 189 | qed | 
| 190 | ||
| 63329 | 191 | lemma interval_lebesgue_integral_add [intro, simp]: | 
| 192 | 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 | 193 | assumes "interval_lebesgue_integrable M a b f" "interval_lebesgue_integrable M a b g" | 
| 63329 | 194 | shows "interval_lebesgue_integrable M a b (\<lambda>x. f x + g x)" and | 
| 195 | "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 | 196 | interval_lebesgue_integral M a b f + interval_lebesgue_integral M a b g" | 
| 68096 | 197 | 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 | 198 | field_simps) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 199 | |
| 63329 | 200 | lemma interval_lebesgue_integral_diff [intro, simp]: | 
| 201 | 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 | 202 | 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 | 203 | "interval_lebesgue_integrable M a b g" | 
| 63329 | 204 | shows "interval_lebesgue_integrable M a b (\<lambda>x. f x - g x)" and | 
| 205 | "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 | 206 | interval_lebesgue_integral M a b f - interval_lebesgue_integral M a b g" | 
| 68096 | 207 | 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 | 208 | field_simps) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 209 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 210 | 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 | 211 |   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 | 212 | 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 | 213 | 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 | 214 | 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 | 215 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 216 | 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 | 217 |   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 | 218 | 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 | 219 | interval_lebesgue_integrable M a b (\<lambda>x. f x * c)" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 220 | 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 | 221 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 222 | lemma interval_lebesgue_integrable_divide [intro, simp]: | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59587diff
changeset | 223 |   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 | 224 | 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 | 225 | interval_lebesgue_integrable M a b (\<lambda>x. f x / c)" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 226 | 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 | 227 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 228 | 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 | 229 |   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 | 230 | 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 | 231 | 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 | 232 | 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 | 233 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 234 | 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 | 235 |   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 | 236 | 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 | 237 | interval_lebesgue_integral M a b f * c" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 238 | 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 | 239 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 240 | lemma interval_lebesgue_integral_divide [simp]: | 
| 59867 
58043346ca64
given up separate type classes demanding `inverse 0 = 0`
 haftmann parents: 
59587diff
changeset | 241 |   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 | 242 | 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 | 243 | 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 | 244 | 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 | 245 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 246 | lemma interval_lebesgue_integral_uminus: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 247 | "interval_lebesgue_integral M a b (\<lambda>x. - f x) = - interval_lebesgue_integral M a b f" | 
| 68096 | 248 | 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 | 249 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 250 | 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 | 251 | "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 | 252 | 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 | 253 | unfolding interval_lebesgue_integral_def | 
| 68096 | 254 | 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 | 255 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 256 | 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 | 257 | fixes a b f | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 258 | assumes "a \<le> b" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 259 | 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: 
69680diff
changeset | 260 | 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 | 261 | |
| 63329 | 262 | 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 | 263 | fixes a b f | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 264 | assumes "a > b" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 265 | shows "interval_lebesgue_integral M a b f = -(LINT x : einterval b a | M. f x)" | 
| 68096 | 266 | 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 | 267 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 268 | 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 | 269 | fixes a b f | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 270 | assumes "a > b" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 271 | shows "interval_lebesgue_integral M a b f = - interval_lebesgue_integral M b a f" | 
| 68096 | 272 | 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 | 273 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 274 | 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: 
66164diff
changeset | 275 | 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 | 276 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 277 | 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: 
66164diff
changeset | 278 | 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 | 279 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 280 | lemma interval_integrable_endpoints_reverse: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 281 | "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 | 282 | 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 | 283 | 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 | 284 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 285 | lemma interval_integral_reflect: | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 286 | "(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: 
69680diff
changeset | 287 | 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 | 288 | case (sym a b) then show ?case | 
| 68096 | 289 | by (auto simp: interval_lebesgue_integral_def interval_integrable_endpoints_reverse | 
| 62390 | 290 | split: if_split_asm) | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 291 | next | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 292 | case (le a b) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 293 |   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: 
66164diff
changeset | 294 | unfolding interval_lebesgue_integrable_def set_lebesgue_integral_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 295 | apply (rule Bochner_Integration.integral_cong [OF refl]) | 
| 68046 | 296 | by (auto simp: einterval_iff ereal_uminus_le_reorder ereal_uminus_less_reorder not_less | 
| 68403 | 297 | 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 | 298 | split: split_indicator) | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 299 | then show ?case | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 300 | unfolding interval_lebesgue_integral_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 301 | 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 | 302 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 303 | |
| 61897 
bc0fc5499085
Bochner integral: prove dominated convergence at_top
 hoelzl parents: 
61882diff
changeset | 304 | lemma interval_lebesgue_integral_0_infty: | 
| 
bc0fc5499085
Bochner integral: prove dominated convergence at_top
 hoelzl parents: 
61882diff
changeset | 305 |   "interval_lebesgue_integrable M 0 \<infinity> f \<longleftrightarrow> set_integrable M {0<..} f"
 | 
| 
bc0fc5499085
Bochner integral: prove dominated convergence at_top
 hoelzl parents: 
61882diff
changeset | 306 |   "interval_lebesgue_integral M 0 \<infinity> f = (LINT x:{0<..}|M. f x)"
 | 
| 63329 | 307 | unfolding zero_ereal_def | 
| 61897 
bc0fc5499085
Bochner integral: prove dominated convergence at_top
 hoelzl parents: 
61882diff
changeset | 308 | by (auto simp: interval_lebesgue_integral_le_eq interval_lebesgue_integrable_def) | 
| 
bc0fc5499085
Bochner integral: prove dominated convergence at_top
 hoelzl parents: 
61882diff
changeset | 309 | |
| 
bc0fc5499085
Bochner integral: prove dominated convergence at_top
 hoelzl parents: 
61882diff
changeset | 310 | 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: 
61882diff
changeset | 311 | unfolding interval_lebesgue_integral_def by auto | 
| 
bc0fc5499085
Bochner integral: prove dominated convergence at_top
 hoelzl parents: 
61882diff
changeset | 312 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 313 | proposition interval_integrable_to_infinity_eq: "(interval_lebesgue_integrable M a \<infinity> f) = | 
| 61897 
bc0fc5499085
Bochner integral: prove dominated convergence at_top
 hoelzl parents: 
61882diff
changeset | 314 |   (set_integrable M {a<..} f)"
 | 
| 70136 | 315 | unfolding interval_lebesgue_integrable_def by auto | 
| 61897 
bc0fc5499085
Bochner integral: prove dominated convergence at_top
 hoelzl parents: 
61882diff
changeset | 316 | |
| 69683 | 317 | 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 | 318 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 319 | lemma interval_integral_zero [simp]: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 320 | fixes a b :: ereal | 
| 68096 | 321 | shows "LBINT x=a..b. 0 = 0" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 322 | 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 | 323 | by simp | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 324 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 325 | 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 | 326 | fixes a b c :: real | 
| 63329 | 327 | 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: 
66164diff
changeset | 328 | unfolding interval_lebesgue_integral_def interval_lebesgue_integrable_def einterval_eq | 
| 68096 | 329 | 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 | 330 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 331 | 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 | 332 | 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 | 333 | 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 | 334 | 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 | 335 | using assms | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 336 | 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 | 337 | 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 | 338 | 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 | 339 | next | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 340 | 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 | 341 | 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 | 342 | 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 | 343 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 344 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 345 | lemma interval_integral_cong: | 
| 63329 | 346 | 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 | 347 | 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 | 348 | using assms | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 349 | 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 | 350 | 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 | 351 | 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 | 352 | next | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 353 | 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 | 354 | 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 | 355 | intro!: set_lebesgue_integral_cong) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 356 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 357 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 358 | lemma interval_lebesgue_integrable_cong_AE: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 359 | "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 | 360 | 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 | 361 | interval_lebesgue_integrable lborel a b f = interval_lebesgue_integrable lborel a b g" | 
| 68096 | 362 | 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 | 363 | 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 | 364 | 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 | 365 | done | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 366 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 367 | lemma interval_integrable_abs_iff: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 368 | fixes f :: "real \<Rightarrow> real" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 369 | 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 | 370 | 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 | 371 | unfolding interval_lebesgue_integrable_def | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 372 | 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 | 373 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 374 | lemma interval_integral_Icc: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 375 | fixes a b :: real | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 376 |   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 | 377 |   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 | 378 | 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 | 379 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 380 | lemma interval_integral_Icc': | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 381 |   "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: 
59867diff
changeset | 382 |   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 | 383 | 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 | 384 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 385 | lemma interval_integral_Ioc: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 386 |   "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 | 387 |   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 | 388 | 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 | 389 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 390 | (* 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 | 391 | lemma interval_integral_Ioc': | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 392 |   "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: 
59867diff
changeset | 393 |   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 | 394 | 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 | 395 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 396 | lemma interval_integral_Ico: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 397 |   "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 | 398 |   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 | 399 | 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 | 400 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 401 | lemma interval_integral_Ioi: | 
| 61882 | 402 |   "\<bar>a\<bar> < \<infinity> \<Longrightarrow> (LBINT x=a..\<infinity>. f x) = (LBINT x : {real_of_ereal a <..}. f x)"
 | 
| 68096 | 403 | 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 | 404 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 405 | lemma interval_integral_Ioo: | 
| 61882 | 406 |   "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 | 407 | 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 | 408 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 409 | lemma interval_integral_discrete_difference: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 410 |   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 | 411 | assumes "countable X" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 412 | 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 | 413 | 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 | 414 |   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 | 415 | 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: 
66164diff
changeset | 416 | 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 | 417 | 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 | 418 | 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 | 419 | done | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 420 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 421 | lemma interval_integral_sum: | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 422 | fixes a b c :: ereal | 
| 63329 | 423 | 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 | 424 | 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: 
69680diff
changeset | 425 | proof - | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 426 | 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 | 427 |   { 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 | 428 | 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 | 429 | 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 | 430 | 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: 
66164diff
changeset | 431 | 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 | 432 | 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 | 433 | 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 | 434 | 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 | 435 | 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: 
59867diff
changeset | 436 | 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 | 437 | 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: 
66164diff
changeset | 438 | 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: 
66164diff
changeset | 439 | unfolding set_borel_measurable_def | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 440 | 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: 
66164diff
changeset | 441 | qed | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 442 | 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 | 443 | using ord | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 444 | 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 | 445 | 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 | 446 | 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 | 447 | } note 1 = this | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 448 |   { 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 | 449 | 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 | 450 | by (metis add.commute) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 451 | } note 2 = this | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 452 | 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 | 453 | 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 | 454 | show ?thesis | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 455 | using integrable | 
| 73526 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
71827diff
changeset | 456 | apply (cases a b b c a c rule: linorder_le_cases[case_product linorder_le_cases linorder_cases]) | 
| 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
71827diff
changeset | 457 | apply simp_all (* remove some looping cases *) | 
| 
a3cc9fa1295d
new automatic order prover: stateless, complete, verified
 nipkow parents: 
71827diff
changeset | 458 | by (simp_all add: min_absorb1 min_absorb2 max_absorb1 max_absorb2 field_simps 1 2 3) | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 459 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 460 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 461 | lemma interval_integrable_isCont: | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 462 |   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 | 463 | 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 | 464 | interval_lebesgue_integrable lborel a b f" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 465 | 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 | 466 | case (le a b) then show ?case | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 467 | 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 | 468 | by (auto simp: interval_lebesgue_integrable_def | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 469 |         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: 
66164diff
changeset | 470 | 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 | 471 | 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 | 472 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 473 | lemma interval_integrable_continuous_on: | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 474 | 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 | 475 |   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 | 476 | 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 | 477 | 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 | 478 | 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 | 479 | |
| 63329 | 480 | 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 | 481 | 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 | 482 |   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 | 483 | 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 | 484 | |
| 63329 | 485 | 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 | 486 | 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 | 487 | 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 | 488 | by (subst interval_lebesgue_integral_le_eq, simp) (rule set_borel_integral_eq_integral) | 
| 63329 | 489 | |
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 490 | |
| 69683 | 491 | 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 | 492 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 493 | 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 | 494 | fixes a b :: ereal | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 495 | assumes "a < b" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 496 | 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 | 497 |   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 | 498 | "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b" | 
| 61969 | 499 | "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 | 500 | fixes f :: "real \<Rightarrow> real" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 501 |   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 | 502 | 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 | 503 | assumes f_measurable: "set_borel_measurable lborel (einterval a b) f" | 
| 61969 | 504 | assumes lbint_lim: "(\<lambda>i. LBINT x=l i.. u i. f x) \<longlonglongrightarrow> C" | 
| 63329 | 505 | shows | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 506 | "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 | 507 | "(LBINT x=a..b. f x) = C" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 508 | proof - | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 509 |   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 | 510 |   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 | 511 | proof - | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 512 | 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 | 513 | by eventually_elim | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 514 | (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 | 515 | then show ?thesis | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 516 | apply eventually_elim | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 517 | 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 | 518 | 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 | 519 | 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 | 520 | done | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 521 | qed | 
| 61969 | 522 |   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 | 523 | proof - | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 524 |     { 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 | 525 | 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 | 526 | 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 | 527 | 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 | 528 | 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 | 529 | done | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 530 |       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 | 531 | by eventually_elim auto } | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 532 | then show ?thesis | 
| 70365 
4df0628e8545
a few new lemmas and a bit of tidying
 paulson <lp15@cam.ac.uk> parents: 
70136diff
changeset | 533 | 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 | 534 | qed | 
| 61969 | 535 |   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: 
66164diff
changeset | 536 | 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: 
66164diff
changeset | 537 | 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: 
66164diff
changeset | 538 | 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 | 539 | 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: 
66164diff
changeset | 540 | using assms by (simp add: interval_lebesgue_integral_def set_lebesgue_integral_def less_imp_le) | 
| 68096 | 541 | also have "\<dots> = C" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 542 | 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 | 543 | finally show "(LBINT x=a..b. f x) = C" . | 
| 63329 | 544 | show "set_integrable lborel (einterval a b) f" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 545 | unfolding set_integrable_def | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 546 | 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 | 547 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 548 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 549 | 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 | 550 | 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 | 551 |   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 | 552 | assumes "a < b" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 553 |   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 | 554 | "incseq u" "decseq l" "\<And>i. l i < u i" "\<And>i. a < l i" "\<And>i. u i < b" | 
| 61969 | 555 | "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 | 556 | assumes f_integrable: "set_integrable lborel (einterval a b) f" | 
| 61969 | 557 | 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: 
69680diff
changeset | 558 | proof - | 
| 61969 | 559 |   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: 
66164diff
changeset | 560 | 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 | 561 | proof (rule integral_dominated_convergence) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 562 | 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: 
66164diff
changeset | 563 | 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: 
66164diff
changeset | 564 | 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: 
66164diff
changeset | 565 | using f_integrable by (simp add: set_integrable_def) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 566 |     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: 
66164diff
changeset | 567 | 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 | 568 |     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 | 569 | by (intro AE_I2) (auto simp: approx split: split_indicator) | 
| 61969 | 570 |     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: 
70136diff
changeset | 571 | 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 | 572 | fix x | 
| 63329 | 573 |       { fix i assume "l i \<le> x" "x \<le> u i"
 | 
| 61808 | 574 | 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 | 575 | 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 | 576 | 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 | 577 |       then show "eventually (\<lambda>xa. indicator {l xa..u xa} x = (indicator (einterval a b) x::real)) sequentially"
 | 
| 61969 | 578 | 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 | 579 | by (auto split: split_indicator) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 580 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 581 | qed | 
| 61808 | 582 | 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 | 583 | 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 | 584 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 585 | |
| 69683 | 586 | 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: 
66164diff
changeset | 587 | |
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 588 | 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: 
66164diff
changeset | 589 | arbitrary intervals.\<close> | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 590 | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 591 | (* | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 592 | 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 | 593 | *) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 594 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 595 | 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 | 596 | 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 | 597 |   assumes f: "continuous_on {min a b..max a b} f"
 | 
| 63329 | 598 | 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 | 
| 599 |     {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 | 600 | shows "(LBINT x=a..b. f x) = F b - F a" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 601 | proof (cases "a \<le> b") | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 602 | case True | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 603 |   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: 
66164diff
changeset | 604 | by (simp add: True interval_integral_Icc set_lebesgue_integral_def) | 
| 68096 | 605 | also have "\<dots> = F b - F a" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 606 | proof (rule integral_FTC_atLeastAtMost [OF True]) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 607 |     show "continuous_on {a..b} f"
 | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 608 | using True f by linarith | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 609 |     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: 
66164diff
changeset | 610 | 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: 
66164diff
changeset | 611 | qed | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 612 | finally show ?thesis . | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 613 | next | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 614 | case False | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 615 | then have "b \<le> a" | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 616 | by simp | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 617 |   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: 
66164diff
changeset | 618 | by (simp add: \<open>b \<le> a\<close> interval_integral_Icc set_lebesgue_integral_def) | 
| 68096 | 619 | also have "\<dots> = F b - F a" | 
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 620 | 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: 
66164diff
changeset | 621 |     show "continuous_on {b..a} f"
 | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 622 | using False f by linarith | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 623 | 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: 
66164diff
changeset | 624 |          \<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: 
66164diff
changeset | 625 | by (metis F False max_def min_def) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 626 | qed auto | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 627 | finally show ?thesis | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 628 | by (metis interval_integral_endpoints_reverse) | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 629 | qed | 
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 630 | |
| 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 631 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 632 | 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 | 633 | 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 | 634 | assumes "a < b" | 
| 63329 | 635 | assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> DERIV F x :> f x" | 
| 636 | 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 | 637 | assumes f_nonneg: "AE x in lborel. a < ereal x \<longrightarrow> ereal x < b \<longrightarrow> 0 \<le> f x" | 
| 61973 | 638 | assumes A: "((F \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)" | 
| 639 | 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 | 640 | shows | 
| 63329 | 641 | "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 | 642 | "(LBINT x=a..b. f x) = B - A" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 643 | proof - | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 644 | obtain u l where approx: | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 645 |     "einterval a b = (\<Union>i. {l i .. u i})"
 | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 646 | "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: 
68046diff
changeset | 647 | "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b" | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 648 | 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 | 649 | 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 | 650 | 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 | 651 | 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 | 652 | 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 | 653 | 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 | 654 | using assms approx apply (intro interval_integral_FTC_finite) | 
| 68096 | 655 | apply (auto simp: less_imp_le min_def max_def | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 656 | has_real_derivative_iff_has_vector_derivative[symmetric]) | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 657 | 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 | 658 | 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 | 659 |   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 | 660 | proof - | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 661 |     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: 
66164diff
changeset | 662 | 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 | 663 | by (intro borel_integrable_compact f continuous_at_imp_continuous_on compact_Icc ballI) | 
| 68403 | 664 | (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 | 665 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 666 | 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: 
66164diff
changeset | 667 | unfolding set_borel_measurable_def | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
63941diff
changeset | 668 | 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 | 669 | simp: continuous_on_eq_continuous_at einterval_iff f) | 
| 61969 | 670 | 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 | 671 | apply (subst FTCi) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 672 | apply (intro tendsto_intros) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 673 | 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 | 674 | 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 | 675 | 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 | 676 | 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 | 677 | 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 | 678 | show "(LBINT x=a..b. f x) = B - A" | 
| 61808 | 679 | by (rule interval_integral_Icc_approx_nonneg [OF \<open>a < b\<close> approx 1 f_nonneg 2 3]) | 
| 63329 | 680 | show "set_integrable lborel (einterval a b) f" | 
| 61808 | 681 | 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 | 682 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 683 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 684 | 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 | 685 | 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 | 686 | assumes "a < b" | 
| 63329 | 687 | assumes F: "\<And>x. a < ereal x \<Longrightarrow> ereal x < b \<Longrightarrow> (F has_vector_derivative f x) (at x)" | 
| 688 | 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 | 689 | assumes f_integrable: "set_integrable lborel (einterval a b) f" | 
| 61973 | 690 | assumes A: "((F \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)" | 
| 691 | 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 | 692 | shows "(LBINT x=a..b. f x) = B - A" | 
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 693 | proof - | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 694 | obtain u l where approx: | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 695 |     "einterval a b = (\<Union>i. {l i .. u i})"
 | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 696 | "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: 
68046diff
changeset | 697 | "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b" | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 698 | 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 | 699 | 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 | 700 | 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 | 701 | 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 | 702 | 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 | 703 | 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 | 704 | using assms approx | 
| 68096 | 705 | 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 | 706 | 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 | 707 | intro: has_vector_derivative_at_within) | 
| 61969 | 708 | have "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> B - A" | 
| 68096 | 709 | unfolding FTCi | 
| 710 | proof (intro tendsto_intros) | |
| 711 | show "(\<lambda>x. F (l x)) \<longlonglongrightarrow> A" | |
| 712 | using A approx unfolding tendsto_at_iff_sequentially comp_def | |
| 713 | by (elim allE[of _ "\<lambda>i. ereal (l i)"], auto) | |
| 714 | show "(\<lambda>x. F (u x)) \<longlonglongrightarrow> B" | |
| 715 | using B approx unfolding tendsto_at_iff_sequentially comp_def | |
| 716 | by (elim allE[of _ "\<lambda>i. ereal (u i)"], auto) | |
| 717 | qed | |
| 61969 | 718 | moreover have "(\<lambda>i. LBINT x=l i..u i. f x) \<longlonglongrightarrow> (LBINT x=a..b. f x)" | 
| 61808 | 719 | 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 | 720 | ultimately show ?thesis | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 721 | by (elim LIMSEQ_unique) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 722 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 723 | |
| 63329 | 724 | (* | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 725 | 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 | 726 | einterval. | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 727 | *) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 728 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 729 | theorem interval_integral_FTC2: | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 730 | 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 | 731 | 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 | 732 |   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 | 733 | fixes x :: real | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 734 | 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 | 735 |   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: 
69680diff
changeset | 736 | proof - | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 737 | 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 | 738 |   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 | 739 | 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 | 740 |   have "((\<lambda>u. integral {a..u} f) has_vector_derivative f x) (at x within {a..b})"
 | 
| 68096 | 741 | using \<open>a \<le> x\<close> \<open>x \<le> b\<close> | 
| 742 | 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 | 743 |   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 | 744 | by simp | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 745 |   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 | 746 | 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 | 747 | (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 | 748 |   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 | 749 | 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 | 750 | then show ?thesis | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 751 | 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 | 752 |     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 | 753 | 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 | 754 | using assms | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 755 | apply (intro interval_integral_sum) | 
| 68096 | 756 | apply (auto simp: interval_lebesgue_integrable_def simp del: real_scaleR_def) | 
| 757 | 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 | 758 | qed (insert assms, auto) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 759 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 760 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 761 | proposition einterval_antiderivative: | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 762 | 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 | 763 | 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 | 764 | 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: 
69680diff
changeset | 765 | proof - | 
| 63329 | 766 | from einterval_nonempty [OF \<open>a < b\<close>] obtain c :: real where [simp]: "a < c" "c < b" | 
| 68096 | 767 | 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 | 768 | 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 | 769 | show ?thesis | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 770 | proof (rule exI, clarsimp) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 771 | fix x :: real | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 772 | 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 | 773 | have 1: "a < min c x" by simp | 
| 63329 | 774 | from einterval_nonempty [OF 1] obtain d :: real where [simp]: "a < d" "d < c" "d < x" | 
| 68096 | 775 | 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 | 776 | have 2: "max c x < b" by simp | 
| 63329 | 777 | from einterval_nonempty [OF 2] obtain e :: real where [simp]: "c < e" "x < e" "e < b" | 
| 68096 | 778 | by (auto simp: einterval_def) | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 779 |     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: 
68046diff
changeset | 780 |     proof (rule has_vector_derivative_within_subset [of _ _ _ "{d..e}"])
 | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 781 |       have "continuous_on {d..e} f"
 | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 782 | proof (intro continuous_at_imp_continuous_on ballI contf; clarsimp) | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 783 | 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: 
68046diff
changeset | 784 | 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: 
68046diff
changeset | 785 | 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: 
68046diff
changeset | 786 | 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: 
68046diff
changeset | 787 | qed | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 788 |       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: 
68046diff
changeset | 789 | 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: 
68046diff
changeset | 790 | qed auto | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 791 | then show "(?F has_vector_derivative f x) (at x)" | 
| 68096 | 792 |       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 | 793 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 794 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 795 | |
| 69683 | 796 | 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 | 797 | |
| 67974 
3f352a91b45a
replacement of set integral abbreviations by actual definitions!
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 798 | 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: 
66164diff
changeset | 799 | arbitrary intervals.\<close> | 
| 63329 | 800 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 801 | 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 | 802 | 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 | 803 | assumes "a \<le> b" | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 804 |   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 | 805 |   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 | 806 |   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 | 807 | 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: 
69680diff
changeset | 808 | proof- | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 809 |   have v_derivg: "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> (g has_vector_derivative (g' x)) (at x within {a..b})"
 | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 810 | using derivg unfolding has_real_derivative_iff_has_vector_derivative . | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 811 |   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 | 812 | by (rule continuous_on_vector_derivative) auto | 
| 68096 | 813 |   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
 | 
| 814 | 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>) | |
| 815 |   obtain c d where g_im: "g ` {a..b} = {c..d}" and "c \<le> d"
 | |
| 816 | by (metis continuous_image_closed_interval contg \<open>a \<le> b\<close>) | |
| 817 | obtain F where derivF: | |
| 818 |          "\<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}))" 
 | |
| 819 | using continuous_on_subset [OF contf] g_im | |
| 820 | 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 | 821 |   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 | 822 | by (blast intro: continuous_on_compose2 contf contg) | 
| 68096 | 823 |   have "LBINT x. indicat_real {a..b} x *\<^sub>R g' x *\<^sub>R f (g x) = F (g b) - F (g a)"
 | 
| 824 | apply (rule integral_FTC_atLeastAtMost | |
| 825 | [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 | 826 | 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 | 827 | done | 
| 68096 | 828 | then have "LBINT x=a..b. g' x *\<^sub>R f (g x) = F (g b) - F (g a)" | 
| 829 | 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 | 830 | moreover have "LBINT y=(g a)..(g b). f y = F (g b) - F (g a)" | 
| 68096 | 831 | proof (rule interval_integral_FTC_finite) | 
| 832 |     show "continuous_on {min (g a) (g b)..max (g a) (g b)} f"
 | |
| 833 | by (rule continuous_on_subset [OF contf]) (auto simp: image_def 1) | |
| 834 |     show "(F has_vector_derivative f y) (at y within {min (g a) (g b)..max (g a) (g b)})" 
 | |
| 835 | if y: "min (g a) (g b) \<le> y" "y \<le> max (g a) (g b)" for y | |
| 836 | proof - | |
| 837 | obtain x where "a \<le> x" "x \<le> b" "y = g x" | |
| 838 | using 1 y by force | |
| 839 | then show ?thesis | |
| 840 | by (auto simp: image_def intro!: 1 has_vector_derivative_within_subset [OF derivF]) | |
| 841 | qed | |
| 842 | qed | |
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 843 | ultimately show ?thesis by simp | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 844 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 845 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 846 | (* 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 | 847 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 848 | 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 | 849 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" and a b u v :: ereal | 
| 63329 | 850 | assumes "a < b" | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 851 | 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 | 852 | 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 | 853 | 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 | 854 | and g'_nonneg: "\<And>x. a \<le> ereal x \<Longrightarrow> ereal x \<le> b \<Longrightarrow> 0 \<le> g' x" | 
| 61973 | 855 | and A: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)" | 
| 856 | 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 | 857 | 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 | 858 | 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 | 859 | 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: 
69680diff
changeset | 860 | proof - | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 861 | obtain u l where approx [simp]: | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 862 |     "einterval a b = (\<Union>i. {l i .. u i})"
 | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 863 | "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: 
68046diff
changeset | 864 | "l \<longlonglongrightarrow> a" "u \<longlonglongrightarrow> b" | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 865 | 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 | 866 | note less_imp_le [simp] | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 867 | 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 | 868 | 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 | 869 | 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 | 870 | 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: 
68046diff
changeset | 871 | then have lessb[simp]: "\<And>i. l i < b" | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 872 | using approx(4) less_eq_real_def by blast | 
| 63329 | 873 | 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 | 874 | 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: 
68046diff
changeset | 875 | 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 | 876 | 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: 
68046diff
changeset | 877 | 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: 
68532diff
changeset | 878 | 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: 
68532diff
changeset | 879 | 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: 
68046diff
changeset | 880 | 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: 
68532diff
changeset | 881 | 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: 
68046diff
changeset | 882 | 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: 
68046diff
changeset | 883 | qed | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 884 |   have "A \<le> B" and un: "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
 | 
| 63329 | 885 | proof - | 
| 61969 | 886 | have A2: "(\<lambda>i. g (l i)) \<longlonglongrightarrow> A" | 
| 68096 | 887 | 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 | 888 | 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 | 889 | 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: 
68403diff
changeset | 890 | by (intro decseq_ge, auto simp: decseq_def) | 
| 61969 | 891 | have B2: "(\<lambda>i. g (u i)) \<longlonglongrightarrow> B" | 
| 68096 | 892 | 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 | 893 | 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 | 894 | hence B3: "\<And>i. g (u i) \<le> B" | 
| 68096 | 895 | by (intro incseq_le, auto simp: incseq_def) | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 896 | 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 | 897 | by auto | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 898 | then show "A \<le> B" | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 899 | 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 | 900 |     { fix x :: real
 | 
| 63329 | 901 | 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 | 902 | 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: 
68046diff
changeset | 903 | 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 | 904 | 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 | 905 | 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 | 906 | } note AB = this | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 907 |     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: 
68046diff
changeset | 908 | proof | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 909 |       show "einterval A B \<subseteq> (\<Union>i. {g(l i)<..<g(u i)})"
 | 
| 68096 | 910 | by (auto simp: einterval_def AB) | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 911 |       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: 
68046diff
changeset | 912 | proof (clarsimp simp add: einterval_def, intro conjI) | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 913 | 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: 
68046diff
changeset | 914 | using A3 le_ereal_less by blast | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 915 | 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: 
68046diff
changeset | 916 | using B3 ereal_le_less by blast | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 917 | qed | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 918 | qed | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 919 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 920 | (* finally, the main argument *) | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 921 | 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: 
68046diff
changeset | 922 | apply (rule interval_integral_substitution_finite [OF _ DERIV_subset [OF deriv_g]]) | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 923 | unfolding has_real_derivative_iff_has_vector_derivative[symmetric] | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 924 | apply (auto intro!: continuous_at_imp_continuous_on contf contg') | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 925 | done | 
| 61969 | 926 | 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 | 927 | 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 | 928 | by (rule assms) | 
| 61969 | 929 | 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 | 930 | by (simp add: eq1) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 931 |   have incseq: "incseq (\<lambda>i. {g (l i)<..<g (u i)})"
 | 
| 68096 | 932 | apply (auto simp: incseq_def) | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 933 | 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: 
68046diff
changeset | 934 | by (force intro: less_le_trans) | 
| 68096 | 935 |   have "(\<lambda>i. set_lebesgue_integral lborel {g (l i)<..<g (u i)} f)
 | 
| 936 | \<longlonglongrightarrow> set_lebesgue_integral lborel (einterval A B) f" | |
| 937 | unfolding un by (rule set_integral_cont_up) (use incseq integrable2 un in auto) | |
| 938 | then have "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> (LBINT x = A..B. f x)" | |
| 939 | 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 | 940 | 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 | 941 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 942 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 943 | (* 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 | 944 | 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 | 945 | 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 | 946 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 947 | 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 | 948 | fixes f g g':: "real \<Rightarrow> real" and a b u v :: ereal | 
| 63329 | 949 | assumes "a < b" | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 950 | 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 | 951 | 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 | 952 | 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 | 953 | 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 | 954 | and g'_nonneg: "\<And>x. a \<le> ereal x \<Longrightarrow> ereal x \<le> b \<Longrightarrow> 0 \<le> g' x" | 
| 61973 | 955 | and A: "((ereal \<circ> g \<circ> real_of_ereal) \<longlongrightarrow> A) (at_right a)" | 
| 956 | 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 | 957 | and integrable_fg: "set_integrable lborel (einterval a b) (\<lambda>x. f (g x) * g' x)" | 
| 63329 | 958 | shows | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 959 | "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 | 960 | "(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: 
69680diff
changeset | 961 | proof - | 
| 74362 | 962 | from einterval_Icc_approximation[OF \<open>a < b\<close>] obtain u l where approx [simp]: | 
| 963 |     "einterval a b = (\<Union>i. {l i..u i})"
 | |
| 964 | "incseq u" | |
| 965 | "decseq l" | |
| 966 | "\<And>i. l i < u i" | |
| 967 | "\<And>i. a < ereal (l i)" | |
| 968 | "\<And>i. ereal (u i) < b" | |
| 969 | "(\<lambda>x. ereal (l x)) \<longlonglongrightarrow> a" | |
| 970 | "(\<lambda>x. ereal (u x)) \<longlonglongrightarrow> b" by this auto | |
| 68096 | 971 | 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 | 972 | by (rule order_less_le_trans, rule approx, force) | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 973 | 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 | 974 | 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: 
68046diff
changeset | 975 | have llb[simp]: "\<And>i. l i < b" | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 976 | using lessb approx(4) less_eq_real_def by blast | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 977 | 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 | 978 | 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 | 979 | 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: 
68046diff
changeset | 980 | 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: 
68046diff
changeset | 981 | 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: 
68532diff
changeset | 982 | 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: 
68532diff
changeset | 983 | 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: 
68046diff
changeset | 984 | 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: 
68532diff
changeset | 985 | 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: 
68046diff
changeset | 986 | 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: 
68046diff
changeset | 987 | qed | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 988 |   have "A \<le> B" and un: "einterval A B = (\<Union>i. {g(l i)<..<g(u i)})"
 | 
| 63329 | 989 | proof - | 
| 61969 | 990 | have A2: "(\<lambda>i. g (l i)) \<longlonglongrightarrow> A" | 
| 68096 | 991 | 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 | 992 | 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 | 993 | 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: 
68403diff
changeset | 994 | by (intro decseq_ge, auto simp: decseq_def) | 
| 61969 | 995 | have B2: "(\<lambda>i. g (u i)) \<longlonglongrightarrow> B" | 
| 68096 | 996 | 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 | 997 | 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 | 998 | hence B3: "\<And>i. g (u i) \<le> B" | 
| 68096 | 999 | by (intro incseq_le, auto simp: incseq_def) | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1000 | have "ereal (g (l 0)) \<le> ereal (g (u 0))" | 
| 74362 | 1001 | by (auto simp: less_imp_le) | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1002 | then show "A \<le> B" | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1003 | 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 | 1004 |     { fix x :: real
 | 
| 63329 | 1005 | 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 | 1006 | then have "eventually (\<lambda>i. ereal (g (l i)) < x \<and> x < ereal (g (u i))) sequentially" | 
| 68096 | 1007 | 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 | 1008 | 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 | 1009 | 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 | 1010 | } note AB = this | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1011 |     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: 
68046diff
changeset | 1012 | proof | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1013 |       show "einterval A B \<subseteq> (\<Union>i. {g (l i)<..<g (u i)})"
 | 
| 68096 | 1014 | by (auto simp: einterval_def AB) | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1015 |       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: 
68046diff
changeset | 1016 | apply (clarsimp simp: einterval_def, intro conjI) | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1017 | using A3 le_ereal_less apply blast | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1018 | using B3 ereal_le_less by blast | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1019 | qed | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1020 | qed | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1021 | (* finally, the main argument *) | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1022 | 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: 
68046diff
changeset | 1023 | proof - | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1024 | 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: 
68046diff
changeset | 1025 | apply (rule interval_integral_substitution_finite [OF _ DERIV_subset [OF deriv_g]]) | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 1026 | unfolding has_real_derivative_iff_has_vector_derivative[symmetric] | 
| 74362 | 1027 | apply (auto simp: less_imp_le intro!: continuous_at_imp_continuous_on contf contg') | 
| 68095 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1028 | done | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1029 | then show ?thesis | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1030 | by (simp add: ac_simps) | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1031 | qed | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1032 |   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: 
68046diff
changeset | 1033 | apply (clarsimp simp add: incseq_def, intro conjI) | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1034 | 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: 
68046diff
changeset | 1035 | 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: 
68046diff
changeset | 1036 | 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: 
68046diff
changeset | 1037 | proof - | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1038 |     have "continuous_on {l i..u i} g"
 | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1039 | 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: 
68046diff
changeset | 1040 | with that show ?thesis | 
| 
4fa3e63ecc7e
starting to tidy up Interval_Integral.thy
 paulson <lp15@cam.ac.uk> parents: 
68046diff
changeset | 1041 | 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: 
68046diff
changeset | 1042 | qed | 
| 68096 | 1043 |   have "continuous_on {g (l i)..g (u i)} f" for i
 | 
| 1044 | apply (intro continuous_intros continuous_at_imp_continuous_on) | |
| 1045 | using contf img by force | |
| 1046 |   then have int_f: "\<And>i. set_integrable lborel {g (l i)<..<g (u i)} f"
 | |
| 1047 | 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 | 1048 |   have integrable: "set_integrable lborel (\<Union>i. {g (l i)<..<g (u i)}) f"
 | 
| 68096 | 1049 | proof (intro pos_integrable_to_top incseq int_f) | 
| 1050 | let ?l = "(LBINT x=a..b. f (g x) * g' x)" | |
| 1051 | have "(\<lambda>i. LBINT x=l i..u i. f (g x) * g' x) \<longlonglongrightarrow> ?l" | |
| 1052 | by (intro assms interval_integral_Icc_approx_integrable [OF \<open>a < b\<close> approx]) | |
| 1053 | hence "(\<lambda>i. (LBINT y=g (l i)..g (u i). f y)) \<longlonglongrightarrow> ?l" | |
| 1054 | by (simp add: eq1) | |
| 1055 |     then show "(\<lambda>i. set_lebesgue_integral lborel {g (l i)<..<g (u i)} f) \<longlonglongrightarrow> ?l"
 | |
| 74362 | 1056 | unfolding interval_lebesgue_integral_def by (auto simp: less_imp_le) | 
| 68096 | 1057 | have "\<And>x i. g (l i) \<le> x \<Longrightarrow> x \<le> g (u i) \<Longrightarrow> 0 \<le> f x" | 
| 1058 | using aless f_nonneg img lessb by blast | |
| 1059 |     then show "\<And>x i. x \<in> {g (l i)<..<g (u i)} \<Longrightarrow> 0 \<le> f x"
 | |
| 1060 | using less_eq_real_def by auto | |
| 1061 | 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 | 1062 | 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 | 1063 | by (simp add: un) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1064 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1065 | 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 | 1066 | 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 | 1067 | 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 | 1068 | 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 | 1069 | qed fact+ | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1070 | 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 | 1071 | by (simp add: ac_simps) | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1072 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1073 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1074 | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63886diff
changeset | 1075 | 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: 
63886diff
changeset | 1076 |   ("(2CLBINT _. _)" [0,60] 60)
 | 
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63886diff
changeset | 1077 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63886diff
changeset | 1078 | 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: 
63886diff
changeset | 1079 | |
| 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63886diff
changeset | 1080 | 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: 
63886diff
changeset | 1081 |   ("(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 | 1082 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1083 | translations | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63886diff
changeset | 1084 | "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 | 1085 | |
| 63329 | 1086 | 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 | 1087 | "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 | 1088 | "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 | 1089 | |
| 63329 | 1090 | 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 | 1091 | "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 | 1092 | "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 | 1093 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1094 | syntax | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1095 | "_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 | 1096 |   ("(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 | 1097 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1098 | translations | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1099 | "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 | 1100 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1101 | proposition interval_integral_norm: | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1102 |   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 | 1103 | 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 | 1104 | norm (LBINT t=a..b. f t) \<le> LBINT t=a..b. norm (f t)" | 
| 70136 | 1105 | using integral_norm_bound[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"] | 
| 1106 | 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 | 1107 | |
| 69681 
689997a8a582
redo tagging-related changes from a06b204527e6, 0f4d4a13dc16, and a8faf6f15da7
 immler parents: 
69680diff
changeset | 1108 | proposition interval_integral_norm2: | 
| 63329 | 1109 | "interval_lebesgue_integrable lborel a b f \<Longrightarrow> | 
| 61945 | 1110 | 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: 
69680diff
changeset | 1111 | 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 | 1112 | 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 | 1113 | 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 | 1114 | next | 
| 63329 | 1115 | case (le a b) | 
| 1116 | 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 | 1117 | using integrable_norm[of lborel "\<lambda>x. indicator (einterval a b) x *\<^sub>R f x"] | 
| 68096 | 1118 | 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 | 1119 | 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 | 1120 | then show ?case | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1121 | 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 | 1122 | qed | 
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1123 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1124 | (* 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 | 1125 | 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 | 1126 | apply (intro interval_integral_FTC_finite continuous_intros) | 
| 75462 
7448423e5dba
Renamed the misleading has_field_derivative_iff_has_vector_derivative. Inserted a number of minor lemmas
 paulson <lp15@cam.ac.uk> parents: 
74362diff
changeset | 1127 | by (auto intro!: derivative_eq_intros simp: has_real_derivative_iff_has_vector_derivative[symmetric]) | 
| 59092 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1128 | |
| 
d469103c0737
add integral substitution theorems from Manuel Eberl, Jeremy Avigad, Luke Serafin, and Sudeep Kanav
 hoelzl parents: diff
changeset | 1129 | end |