| author | wenzelm | 
| Thu, 11 Jan 2018 13:48:17 +0100 | |
| changeset 67405 | e9ab4ad7bd15 | 
| parent 67399 | eab6ce8368fa | 
| child 67683 | 817944aeac3f | 
| child 67685 | bdff8bf0a75b | 
| permissions | -rw-r--r-- | 
| 53399 | 1 | (* Author: John Harrison | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 2 | Author: Robert Himmelmann, TU Muenchen (Translation from HOL light) | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 3 | Huge cleanup by LCP | 
| 53399 | 4 | *) | 
| 5 | ||
| 63594 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63593diff
changeset | 6 | section \<open>Henstock-Kurzweil gauge integration in many dimensions.\<close> | 
| 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63593diff
changeset | 7 | |
| 
bd218a9320b5
HOL-Multivariate_Analysis: rename theories for more descriptive names
 hoelzl parents: 
63593diff
changeset | 8 | theory Henstock_Kurzweil_Integration | 
| 41413 
64cd30d6b0b8
explicit file specifications -- avoid secondary load path;
 wenzelm parents: 
40513diff
changeset | 9 | imports | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 10 | Lebesgue_Measure Tagged_Division | 
| 35172 | 11 | begin | 
| 12 | ||
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 13 | lemma norm_diff2: "\<lbrakk>y = y1 + y2; x = x1 + x2; e = e1 + e2; norm(y1 - x1) \<le> e1; norm(y2 - x2) \<le> e2\<rbrakk> | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 14 | \<Longrightarrow> norm(y-x) \<le> e" | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 15 | using norm_triangle_mono [of "y1 - x1" "e1" "y2 - x2" "e2"] | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 16 | by (simp add: add_diff_add) | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 17 | |
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 18 | lemma setcomp_dot1: "{z. P (z \<bullet> (i,0))} = {(x,y). P(x \<bullet> i)}"
 | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 19 | by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 20 | |
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 21 | lemma setcomp_dot2: "{z. P (z \<bullet> (0,i))} = {(x,y). P(y \<bullet> i)}"
 | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 22 | by auto | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 23 | |
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 24 | lemma Sigma_Int_Paircomp1: "(Sigma A B) \<inter> {(x, y). P x} = Sigma (A \<inter> {x. P x}) B"
 | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 25 | by blast | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 26 | |
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 27 | lemma Sigma_Int_Paircomp2: "(Sigma A B) \<inter> {(x, y). P y} = Sigma A (\<lambda>z. B z \<inter> {y. P y})"
 | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 28 | by blast | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 29 | (* END MOVE *) | 
| 59425 | 30 | |
| 60420 | 31 | subsection \<open>Content (length, area, volume...) of an interval.\<close> | 
| 35172 | 32 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 33 | abbreviation content :: "'a::euclidean_space set \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 34 | where "content s \<equiv> measure lborel s" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 35 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 36 | lemma content_cbox_cases: | 
| 64272 | 37 | "content (cbox a b) = (if \<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i then prod (\<lambda>i. b\<bullet>i - a\<bullet>i) Basis else 0)" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 38 | by (simp add: measure_lborel_cbox_eq inner_diff) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 39 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 40 | lemma content_cbox: "\<forall>i\<in>Basis. a\<bullet>i \<le> b\<bullet>i \<Longrightarrow> content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 41 | unfolding content_cbox_cases by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 42 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 43 | lemma content_cbox': "cbox a b \<noteq> {} \<Longrightarrow> content (cbox a b) = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 44 | by (simp add: box_ne_empty inner_diff) | 
| 49970 | 45 | |
| 66296 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 46 | lemma content_cbox_if: "content (cbox a b) = (if cbox a b = {} then 0 else \<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
 | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 47 | by (simp add: content_cbox') | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 48 | |
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 49 | lemma content_division_of: | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 50 | assumes "K \<in> \<D>" "\<D> division_of S" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 51 | shows "content K = (\<Prod>i \<in> Basis. interval_upperbound K \<bullet> i - interval_lowerbound K \<bullet> i)" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 52 | proof - | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 53 | obtain a b where "K = cbox a b" | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 54 | using cbox_division_memE assms by metis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 55 | then show ?thesis | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 56 | using assms by (force simp: division_of_def content_cbox') | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 57 | qed | 
| 
33a47f2d9edc
New theory of Equiintegrability / Continuity of the indefinite integral / improper integration
 paulson <lp15@cam.ac.uk> parents: 
66294diff
changeset | 58 | |
| 53408 | 59 | lemma content_real: "a \<le> b \<Longrightarrow> content {a..b} = b - a"
 | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 60 | by simp | 
| 56188 | 61 | |
| 66402 | 62 | lemma abs_eq_content: "\<bar>y - x\<bar> = (if x\<le>y then content {x..y} else content {y..x})"
 | 
| 61204 | 63 | by (auto simp: content_real) | 
| 64 | ||
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 65 | lemma content_singleton: "content {a} = 0"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 66 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 67 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 68 | lemma content_unit[iff]: "content (cbox 0 (One::'a::euclidean_space)) = 1" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 69 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 70 | |
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 71 | lemma content_pos_le [iff]: "0 \<le> content X" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 72 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 73 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 74 | corollary content_nonneg [simp]: "~ content (cbox a b) < 0" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 75 | using not_le by blast | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 76 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 77 | lemma content_pos_lt: "\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i \<Longrightarrow> 0 < content (cbox a b)" | 
| 64272 | 78 | by (auto simp: less_imp_le inner_diff box_eq_empty intro!: prod_pos) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 79 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 80 | lemma content_eq_0: "content (cbox a b) = 0 \<longleftrightarrow> (\<exists>i\<in>Basis. b\<bullet>i \<le> a\<bullet>i)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 81 | by (auto simp: content_cbox_cases not_le intro: less_imp_le antisym eq_refl) | 
| 56188 | 82 | |
| 83 | lemma content_eq_0_interior: "content (cbox a b) = 0 \<longleftrightarrow> interior(cbox a b) = {}"
 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 84 | unfolding content_eq_0 interior_cbox box_eq_empty by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 85 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 86 | lemma content_pos_lt_eq: "0 < content (cbox a (b::'a::euclidean_space)) \<longleftrightarrow> (\<forall>i\<in>Basis. a\<bullet>i < b\<bullet>i)" | 
| 64272 | 87 | by (auto simp add: content_cbox_cases less_le prod_nonneg) | 
| 49970 | 88 | |
| 53399 | 89 | lemma content_empty [simp]: "content {} = 0"
 | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 90 | by simp | 
| 35172 | 91 | |
| 60762 | 92 | lemma content_real_if [simp]: "content {a..b} = (if a \<le> b then b - a else 0)"
 | 
| 93 | by (simp add: content_real) | |
| 94 | ||
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 95 | lemma content_subset: "cbox a b \<subseteq> cbox c d \<Longrightarrow> content (cbox a b) \<le> content (cbox c d)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 96 | unfolding measure_def | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 97 | by (intro enn2real_mono emeasure_mono) (auto simp: emeasure_lborel_cbox_eq) | 
| 56188 | 98 | |
| 99 | lemma content_lt_nz: "0 < content (cbox a b) \<longleftrightarrow> content (cbox a b) \<noteq> 0" | |
| 44890 
22f665a2e91c
new fastforce replacing fastsimp - less confusing name
 nipkow parents: 
44522diff
changeset | 100 | unfolding content_pos_lt_eq content_eq_0 unfolding not_ex not_le by fastforce | 
| 35172 | 101 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 102 | lemma content_Pair: "content (cbox (a,c) (b,d)) = content (cbox a b) * content (cbox c d)" | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 103 | unfolding measure_lborel_cbox_eq Basis_prod_def | 
| 64272 | 104 | apply (subst prod.union_disjoint) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 105 | apply (auto simp: bex_Un ball_Un) | 
| 64272 | 106 | apply (subst (1 2) prod.reindex_nontrivial) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 107 | apply auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 108 | done | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 109 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 110 | lemma content_cbox_pair_eq0_D: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 111 | "content (cbox (a,c) (b,d)) = 0 \<Longrightarrow> content (cbox a b) = 0 \<or> content (cbox c d) = 0" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 112 | by (simp add: content_Pair) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 113 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 114 | lemma content_0_subset: "content(cbox a b) = 0 \<Longrightarrow> s \<subseteq> cbox a b \<Longrightarrow> content s = 0" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 115 | using emeasure_mono[of s "cbox a b" lborel] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 116 | by (auto simp: measure_def enn2real_eq_0_iff emeasure_lborel_cbox_eq) | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 117 | |
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 118 | lemma content_split: | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 119 | fixes a :: "'a::euclidean_space" | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 120 | assumes "k \<in> Basis" | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 121 |   shows "content (cbox a b) = content(cbox a b \<inter> {x. x\<bullet>k \<le> c}) + content(cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
 | 
| 64911 | 122 | \<comment> \<open>Prove using measure theory\<close> | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 123 | proof cases | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 124 | note simps = interval_split[OF assms] content_cbox_cases | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 125 |   have *: "Basis = insert k (Basis - {k})" "\<And>x. finite (Basis-{x})" "\<And>x. x\<notin>Basis-{x}"
 | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 126 | using assms by auto | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 127 |   have *: "\<And>X Y Z. (\<Prod>i\<in>Basis. Z i (if i = k then X else Y i)) = Z k X * (\<Prod>i\<in>Basis-{k}. Z i (Y i))"
 | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 128 |     "(\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i) = (\<Prod>i\<in>Basis-{k}. b\<bullet>i - a\<bullet>i) * (b\<bullet>k - a\<bullet>k)"
 | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 129 | apply (subst *(1)) | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 130 | defer | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 131 | apply (subst *(1)) | 
| 64272 | 132 | unfolding prod.insert[OF *(2-)] | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 133 | apply auto | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 134 | done | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 135 | assume as: "\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i" | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 136 | moreover | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 137 | have "\<And>x. min (b \<bullet> k) c = max (a \<bullet> k) c \<Longrightarrow> | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 138 | x * (b\<bullet>k - a\<bullet>k) = x * (max (a \<bullet> k) c - a \<bullet> k) + x * (b \<bullet> k - max (a \<bullet> k) c)" | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 139 | by (auto simp add: field_simps) | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 140 | moreover | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 141 | have **: "(\<Prod>i\<in>Basis. ((\<Sum>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) *\<^sub>R i) \<bullet> i - a \<bullet> i)) = | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 142 | (\<Prod>i\<in>Basis. (if i = k then min (b \<bullet> k) c else b \<bullet> i) - a \<bullet> i)" | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 143 | "(\<Prod>i\<in>Basis. b \<bullet> i - ((\<Sum>i\<in>Basis. (if i = k then max (a \<bullet> k) c else a \<bullet> i) *\<^sub>R i) \<bullet> i)) = | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 144 | (\<Prod>i\<in>Basis. b \<bullet> i - (if i = k then max (a \<bullet> k) c else a \<bullet> i))" | 
| 64272 | 145 | by (auto intro!: prod.cong) | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 146 | have "\<not> a \<bullet> k \<le> c \<Longrightarrow> \<not> c \<le> b \<bullet> k \<Longrightarrow> False" | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 147 | unfolding not_le | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 148 | using as[unfolded ,rule_format,of k] assms | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 149 | by auto | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 150 | ultimately show ?thesis | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 151 | using assms | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 152 | unfolding simps ** | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 153 | unfolding *(1)[of "\<lambda>i x. b\<bullet>i - x"] *(1)[of "\<lambda>i x. x - a\<bullet>i"] | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 154 | unfolding *(2) | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 155 | by auto | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 156 | next | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 157 | assume "\<not> (\<forall>i\<in>Basis. a \<bullet> i \<le> b \<bullet> i)" | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 158 |   then have "cbox a b = {}"
 | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 159 | unfolding box_eq_empty by (auto simp: not_le) | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 160 | then show ?thesis | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 161 | by (auto simp: not_le) | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 162 | qed | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 163 | |
| 49970 | 164 | lemma division_of_content_0: | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 165 | assumes "content (cbox a b) = 0" "d division_of (cbox a b)" "K \<in> d" | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 166 | shows "content K = 0" | 
| 49970 | 167 | unfolding forall_in_division[OF assms(2)] | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 168 | by (meson assms content_0_subset division_of_def) | 
| 49970 | 169 | |
| 64267 | 170 | lemma sum_content_null: | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 171 | assumes "content (cbox a b) = 0" | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 172 | and "p tagged_division_of (cbox a b)" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 173 | shows "(\<Sum>(x,K)\<in>p. content K *\<^sub>R f x) = (0::'a::real_normed_vector)" | 
| 64267 | 174 | proof (rule sum.neutral, rule) | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 175 | fix y | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 176 | assume y: "y \<in> p" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 177 | obtain x K where xk: "y = (x, K)" | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 178 | using surj_pair[of y] by blast | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 179 | then obtain c d where k: "K = cbox c d" "K \<subseteq> cbox a b" | 
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 180 | by (metis assms(2) tagged_division_ofD(3) tagged_division_ofD(4) y) | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 181 | have "(\<lambda>(x',K'). content K' *\<^sub>R f x') y = content K *\<^sub>R f x" | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 182 | unfolding xk by auto | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 183 | also have "\<dots> = 0" | 
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 184 | using assms(1) content_0_subset k(2) by auto | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 185 | finally show "(\<lambda>(x, k). content k *\<^sub>R f x) y = 0" . | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 186 | qed | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 187 | |
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 188 | global_interpretation sum_content: operative plus 0 content | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 189 | rewrites "comm_monoid_set.F plus 0 = sum" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 190 | proof - | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 191 | interpret operative plus 0 content | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 192 | by standard (auto simp add: content_split [symmetric] content_eq_0_interior) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 193 | show "operative plus 0 content" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 194 | by standard | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 195 | show "comm_monoid_set.F plus 0 = sum" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 196 | by (simp add: sum_def) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 197 | qed | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 198 | |
| 64267 | 199 | lemma additive_content_division: "d division_of (cbox a b) \<Longrightarrow> sum content d = content (cbox a b)" | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 200 | by (fact sum_content.division) | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 201 | |
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 202 | lemma additive_content_tagged_division: | 
| 64267 | 203 | "d tagged_division_of (cbox a b) \<Longrightarrow> sum (\<lambda>(x,l). content l) d = content (cbox a b)" | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 204 | by (fact sum_content.tagged_division) | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 205 | |
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 206 | lemma subadditive_content_division: | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 207 | assumes "\<D> division_of S" "S \<subseteq> cbox a b" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 208 | shows "sum content \<D> \<le> content(cbox a b)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 209 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 210 | have "\<D> division_of \<Union>\<D>" "\<Union>\<D> \<subseteq> cbox a b" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 211 | using assms by auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 212 | then obtain \<D>' where "\<D> \<subseteq> \<D>'" "\<D>' division_of cbox a b" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 213 | using partial_division_extend_interval by metis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 214 | then have "sum content \<D> \<le> sum content \<D>'" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 215 | using sum_mono2 by blast | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 216 | also have "... \<le> content(cbox a b)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 217 | by (simp add: \<open>\<D>' division_of cbox a b\<close> additive_content_division less_eq_real_def) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 218 | finally show ?thesis . | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 219 | qed | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 220 | |
| 66402 | 221 | lemma content_real_eq_0: "content {a..b::real} = 0 \<longleftrightarrow> a \<ge> b"
 | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 222 | by (metis atLeastatMost_empty_iff2 content_empty content_real diff_self eq_iff le_cases le_iff_diff_le_0) | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 223 | |
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 224 | lemma property_empty_interval: "\<forall>a b. content (cbox a b) = 0 \<longrightarrow> P (cbox a b) \<Longrightarrow> P {}"
 | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 225 | using content_empty unfolding empty_as_interval by auto | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 226 | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 227 | lemma interval_bounds_nz_content [simp]: | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 228 | assumes "content (cbox a b) \<noteq> 0" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 229 | shows "interval_upperbound (cbox a b) = b" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 230 | and "interval_lowerbound (cbox a b) = a" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 231 | by (metis assms content_empty interval_bounds')+ | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 232 | |
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 233 | subsection \<open>Gauge integral\<close> | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 234 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 235 | text \<open>Case distinction to define it first on compact intervals first, then use a limit. This is only | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 236 | much later unified. In Fremlin: Measure Theory, Volume 4I this is generalized using residual sets.\<close> | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 237 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 238 | definition has_integral :: "('n::euclidean_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'b \<Rightarrow> 'n set \<Rightarrow> bool"
 | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 239 | (infixr "has'_integral" 46) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 240 | where "(f has_integral I) s \<longleftrightarrow> | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 241 | (if \<exists>a b. s = cbox a b | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 242 | then ((\<lambda>p. \<Sum>(x,k)\<in>p. content k *\<^sub>R f x) \<longlongrightarrow> I) (division_filter s) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 243 | else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 244 | (\<exists>z. ((\<lambda>p. \<Sum>(x,k)\<in>p. content k *\<^sub>R (if x \<in> s then f x else 0)) \<longlongrightarrow> z) (division_filter (cbox a b)) \<and> | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 245 | norm (z - I) < e)))" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 246 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 247 | lemma has_integral_cbox: | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 248 | "(f has_integral I) (cbox a b) \<longleftrightarrow> ((\<lambda>p. \<Sum>(x,k)\<in>p. content k *\<^sub>R f x) \<longlongrightarrow> I) (division_filter (cbox a b))" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 249 | by (auto simp add: has_integral_def) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 250 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 251 | lemma has_integral: | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 252 | "(f has_integral y) (cbox a b) \<longleftrightarrow> | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 253 | (\<forall>e>0. \<exists>\<gamma>. gauge \<gamma> \<and> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 254 | (\<forall>\<D>. \<D> tagged_division_of (cbox a b) \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 255 | norm (sum (\<lambda>(x,k). content(k) *\<^sub>R f x) \<D> - y) < e))" | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 256 | by (auto simp: dist_norm eventually_division_filter has_integral_def tendsto_iff) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 257 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 258 | lemma has_integral_real: | 
| 66402 | 259 |   "(f has_integral y) {a..b::real} \<longleftrightarrow>
 | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 260 | (\<forall>e>0. \<exists>\<gamma>. gauge \<gamma> \<and> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 261 |       (\<forall>\<D>. \<D> tagged_division_of {a..b} \<and> \<gamma> fine \<D> \<longrightarrow>
 | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 262 | norm (sum (\<lambda>(x,k). content(k) *\<^sub>R f x) \<D> - y) < e))" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 263 | unfolding box_real[symmetric] by (rule has_integral) | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 264 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 265 | lemma has_integralD[dest]: | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 266 | assumes "(f has_integral y) (cbox a b)" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 267 | and "e > 0" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 268 | obtains \<gamma> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 269 | where "gauge \<gamma>" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 270 | and "\<And>\<D>. \<D> tagged_division_of (cbox a b) \<Longrightarrow> \<gamma> fine \<D> \<Longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 271 | norm ((\<Sum>(x,k)\<in>\<D>. content k *\<^sub>R f x) - y) < e" | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 272 | using assms unfolding has_integral by auto | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 273 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 274 | lemma has_integral_alt: | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 275 | "(f has_integral y) i \<longleftrightarrow> | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 276 | (if \<exists>a b. i = cbox a b | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 277 | then (f has_integral y) i | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 278 | else (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 279 | (\<exists>z. ((\<lambda>x. if x \<in> i then f x else 0) has_integral z) (cbox a b) \<and> norm (z - y) < e)))" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 280 | by (subst has_integral_def) (auto simp add: has_integral_cbox) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 281 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 282 | lemma has_integral_altD: | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 283 | assumes "(f has_integral y) i" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 284 | and "\<not> (\<exists>a b. i = cbox a b)" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 285 | and "e>0" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 286 | obtains B where "B > 0" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 287 | and "\<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 288 | (\<exists>z. ((\<lambda>x. if x \<in> i then f(x) else 0) has_integral z) (cbox a b) \<and> norm(z - y) < e)" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 289 | using assms has_integral_alt[of f y i] by auto | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 290 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 291 | definition integrable_on (infixr "integrable'_on" 46) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 292 | where "f integrable_on i \<longleftrightarrow> (\<exists>y. (f has_integral y) i)" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 293 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 294 | definition "integral i f = (SOME y. (f has_integral y) i \<or> ~ f integrable_on i \<and> y=0)" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 295 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 296 | lemma integrable_integral[intro]: "f integrable_on i \<Longrightarrow> (f has_integral (integral i f)) i" | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 297 | unfolding integrable_on_def integral_def by (metis (mono_tags, lifting) someI_ex) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 298 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 299 | lemma not_integrable_integral: "~ f integrable_on i \<Longrightarrow> integral i f = 0" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 300 | unfolding integrable_on_def integral_def by blast | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 301 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 302 | lemma has_integral_integrable[dest]: "(f has_integral i) s \<Longrightarrow> f integrable_on s" | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 303 | unfolding integrable_on_def by auto | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 304 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 305 | lemma has_integral_integral: "f integrable_on s \<longleftrightarrow> (f has_integral (integral s f)) s" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 306 | by auto | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 307 | |
| 60420 | 308 | subsection \<open>Basic theorems about integrals.\<close> | 
| 35172 | 309 | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 310 | lemma has_integral_eq_rhs: "(f has_integral j) S \<Longrightarrow> i = j \<Longrightarrow> (f has_integral i) S" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 311 | by (rule forw_subst) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 312 | |
| 66519 | 313 | lemma has_integral_unique_cbox: | 
| 314 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" | |
| 315 | shows "(f has_integral k1) (cbox a b) \<Longrightarrow> (f has_integral k2) (cbox a b) \<Longrightarrow> k1 = k2" | |
| 316 | by (auto simp: has_integral_cbox intro: tendsto_unique[OF division_filter_not_empty]) | |
| 317 | ||
| 53409 | 318 | lemma has_integral_unique: | 
| 56188 | 319 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" | 
| 66519 | 320 | assumes "(f has_integral k1) i" "(f has_integral k2) i" | 
| 53409 | 321 | shows "k1 = k2" | 
| 53410 | 322 | proof (rule ccontr) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 323 | let ?e = "norm (k1 - k2)/2" | 
| 66519 | 324 | let ?F = "(\<lambda>x. if x \<in> i then f x else 0)" | 
| 325 | assume "k1 \<noteq> k2" | |
| 53410 | 326 | then have e: "?e > 0" | 
| 327 | by auto | |
| 66519 | 328 | have nonbox: "\<not> (\<exists>a b. i = cbox a b)" | 
| 329 | using \<open>k1 \<noteq> k2\<close> assms has_integral_unique_cbox by blast | |
| 55751 | 330 | obtain B1 where B1: | 
| 331 | "0 < B1" | |
| 56188 | 332 | "\<And>a b. ball 0 B1 \<subseteq> cbox a b \<Longrightarrow> | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 333 | \<exists>z. (?F has_integral z) (cbox a b) \<and> norm (z - k1) < norm (k1 - k2)/2" | 
| 66519 | 334 | by (rule has_integral_altD[OF assms(1) nonbox,OF e]) blast | 
| 55751 | 335 | obtain B2 where B2: | 
| 336 | "0 < B2" | |
| 56188 | 337 | "\<And>a b. ball 0 B2 \<subseteq> cbox a b \<Longrightarrow> | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 338 | \<exists>z. (?F has_integral z) (cbox a b) \<and> norm (z - k2) < norm (k1 - k2)/2" | 
| 66519 | 339 | by (rule has_integral_altD[OF assms(2) nonbox,OF e]) blast | 
| 340 | obtain a b :: 'n where ab: "ball 0 B1 \<subseteq> cbox a b" "ball 0 B2 \<subseteq> cbox a b" | |
| 341 | by (metis Un_subset_iff bounded_Un bounded_ball bounded_subset_cbox) | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 342 | obtain w where w: "(?F has_integral w) (cbox a b)" "norm (w - k1) < norm (k1 - k2)/2" | 
| 53410 | 343 | using B1(2)[OF ab(1)] by blast | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 344 | obtain z where z: "(?F has_integral z) (cbox a b)" "norm (z - k2) < norm (k1 - k2)/2" | 
| 53410 | 345 | using B2(2)[OF ab(2)] by blast | 
| 346 | have "z = w" | |
| 66519 | 347 | using has_integral_unique_cbox[OF w(1) z(1)] by auto | 
| 53410 | 348 | then have "norm (k1 - k2) \<le> norm (z - k2) + norm (w - k1)" | 
| 349 | using norm_triangle_ineq4 [of "k1 - w" "k2 - z"] | |
| 350 | by (auto simp add: norm_minus_commute) | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 351 | also have "\<dots> < norm (k1 - k2)/2 + norm (k1 - k2)/2" | 
| 66519 | 352 | by (metis add_strict_mono z(2) w(2)) | 
| 53410 | 353 | finally show False by auto | 
| 354 | qed | |
| 355 | ||
| 356 | lemma integral_unique [intro]: "(f has_integral y) k \<Longrightarrow> integral k f = y" | |
| 357 | unfolding integral_def | |
| 358 | by (rule some_equality) (auto intro: has_integral_unique) | |
| 359 | ||
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 360 | lemma eq_integralD: "integral k f = y \<Longrightarrow> (f has_integral y) k \<or> ~ f integrable_on k \<and> y=0" | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 361 | unfolding integral_def integrable_on_def | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 362 | apply (erule subst) | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 363 | apply (rule someI_ex) | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 364 | by blast | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 365 | |
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 366 | lemma has_integral_const [intro]: | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 367 | fixes a b :: "'a::euclidean_space" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 368 | shows "((\<lambda>x. c) has_integral (content (cbox a b) *\<^sub>R c)) (cbox a b)" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 369 | using eventually_division_filter_tagged_division[of "cbox a b"] | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 370 | additive_content_tagged_division[of _ a b] | 
| 64267 | 371 | by (auto simp: has_integral_cbox split_beta' scaleR_sum_left[symmetric] | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 372 | elim!: eventually_mono intro!: tendsto_cong[THEN iffD1, OF _ tendsto_const]) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 373 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 374 | lemma has_integral_const_real [intro]: | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 375 | fixes a b :: real | 
| 66402 | 376 |   shows "((\<lambda>x. c) has_integral (content {a..b} *\<^sub>R c)) {a..b}"
 | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 377 | by (metis box_real(2) has_integral_const) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 378 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 379 | lemma has_integral_integrable_integral: "(f has_integral i) s \<longleftrightarrow> f integrable_on s \<and> integral s f = i" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 380 | by blast | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 381 | |
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 382 | lemma integral_const [simp]: | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 383 | fixes a b :: "'a::euclidean_space" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 384 | shows "integral (cbox a b) (\<lambda>x. c) = content (cbox a b) *\<^sub>R c" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 385 | by (rule integral_unique) (rule has_integral_const) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 386 | |
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 387 | lemma integral_const_real [simp]: | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 388 | fixes a b :: real | 
| 66402 | 389 |   shows "integral {a..b} (\<lambda>x. c) = content {a..b} *\<^sub>R c"
 | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 390 | by (metis box_real(2) integral_const) | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 391 | |
| 66519 | 392 | lemma has_integral_is_0_cbox: | 
| 56188 | 393 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" | 
| 66519 | 394 | assumes "\<And>x. x \<in> cbox a b \<Longrightarrow> f x = 0" | 
| 395 | shows "(f has_integral 0) (cbox a b)" | |
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 396 | unfolding has_integral_cbox | 
| 66519 | 397 | using eventually_division_filter_tagged_division[of "cbox a b"] assms | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 398 | by (subst tendsto_cong[where g="\<lambda>_. 0"]) | 
| 64267 | 399 | (auto elim!: eventually_mono intro!: sum.neutral simp: tag_in_interval) | 
| 66519 | 400 | |
| 401 | lemma has_integral_is_0: | |
| 402 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" | |
| 403 | assumes "\<And>x. x \<in> S \<Longrightarrow> f x = 0" | |
| 404 | shows "(f has_integral 0) S" | |
| 405 | proof (cases "(\<exists>a b. S = cbox a b)") | |
| 406 | case True with assms has_integral_is_0_cbox show ?thesis | |
| 407 | by blast | |
| 408 | next | |
| 409 | case False | |
| 410 | have *: "(\<lambda>x. if x \<in> S then f x else 0) = (\<lambda>x. 0)" | |
| 411 | by (auto simp add: assms) | |
| 412 | show ?thesis | |
| 413 | using has_integral_is_0_cbox False | |
| 60396 | 414 | by (subst has_integral_alt) (force simp add: *) | 
| 53410 | 415 | qed | 
| 35172 | 416 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 417 | lemma has_integral_0[simp]: "((\<lambda>x::'n::euclidean_space. 0) has_integral 0) S" | 
| 53410 | 418 | by (rule has_integral_is_0) auto | 
| 35172 | 419 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 420 | lemma has_integral_0_eq[simp]: "((\<lambda>x. 0) has_integral i) S \<longleftrightarrow> i = 0" | 
| 35172 | 421 | using has_integral_unique[OF has_integral_0] by auto | 
| 422 | ||
| 66519 | 423 | lemma has_integral_linear_cbox: | 
| 424 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" | |
| 425 | assumes f: "(f has_integral y) (cbox a b)" | |
| 426 | and h: "bounded_linear h" | |
| 427 | shows "((h \<circ> f) has_integral (h y)) (cbox a b)" | |
| 428 | proof - | |
| 429 | interpret bounded_linear h using h . | |
| 430 | show ?thesis | |
| 431 | unfolding has_integral_cbox using tendsto [OF f [unfolded has_integral_cbox]] | |
| 432 | by (simp add: sum scaleR split_beta') | |
| 433 | qed | |
| 434 | ||
| 53410 | 435 | lemma has_integral_linear: | 
| 56188 | 436 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" | 
| 66519 | 437 | assumes f: "(f has_integral y) S" | 
| 438 | and h: "bounded_linear h" | |
| 439 | shows "((h \<circ> f) has_integral (h y)) S" | |
| 440 | proof (cases "(\<exists>a b. S = cbox a b)") | |
| 441 | case True with f h has_integral_linear_cbox show ?thesis | |
| 442 | by blast | |
| 443 | next | |
| 444 | case False | |
| 445 | interpret bounded_linear h using h . | |
| 53410 | 446 | from pos_bounded obtain B where B: "0 < B" "\<And>x. norm (h x) \<le> norm x * B" | 
| 447 | by blast | |
| 66519 | 448 | let ?S = "\<lambda>f x. if x \<in> S then f x else 0" | 
| 449 | show ?thesis | |
| 450 | proof (subst has_integral_alt, clarsimp simp: False) | |
| 53410 | 451 | fix e :: real | 
| 452 | assume e: "e > 0" | |
| 56541 | 453 | have *: "0 < e/B" using e B(1) by simp | 
| 53410 | 454 | obtain M where M: | 
| 455 | "M > 0" | |
| 56188 | 456 | "\<And>a b. ball 0 M \<subseteq> cbox a b \<Longrightarrow> | 
| 66519 | 457 | \<exists>z. (?S f has_integral z) (cbox a b) \<and> norm (z - y) < e/B" | 
| 458 | using has_integral_altD[OF f False *] by blast | |
| 56188 | 459 | show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | 
| 66519 | 460 | (\<exists>z. (?S(h \<circ> f) has_integral z) (cbox a b) \<and> norm (z - h y) < e)" | 
| 461 | proof (rule exI, intro allI conjI impI) | |
| 462 | show "M > 0" using M by metis | |
| 463 | next | |
| 464 | fix a b::'n | |
| 465 | assume sb: "ball 0 M \<subseteq> cbox a b" | |
| 466 | obtain z where z: "(?S f has_integral z) (cbox a b)" "norm (z - y) < e/B" | |
| 467 | using M(2)[OF sb] by blast | |
| 468 | have *: "?S(h \<circ> f) = h \<circ> ?S f" | |
| 60396 | 469 | using zero by auto | 
| 66519 | 470 | show "\<exists>z. (?S(h \<circ> f) has_integral z) (cbox a b) \<and> norm (z - h y) < e" | 
| 53410 | 471 | apply (rule_tac x="h z" in exI) | 
| 66519 | 472 | apply (simp add: * has_integral_linear_cbox[OF z(1) h]) | 
| 61165 | 473 | apply (metis B diff le_less_trans pos_less_divide_eq z(2)) | 
| 474 | done | |
| 53410 | 475 | qed | 
| 476 | qed | |
| 477 | qed | |
| 478 | ||
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 479 | lemma has_integral_scaleR_left: | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 480 | "(f has_integral y) S \<Longrightarrow> ((\<lambda>x. f x *\<^sub>R c) has_integral (y *\<^sub>R c)) S" | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 481 | using has_integral_linear[OF _ bounded_linear_scaleR_left] by (simp add: comp_def) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 482 | |
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 483 | lemma integrable_on_scaleR_left: | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 484 | assumes "f integrable_on A" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 485 | shows "(\<lambda>x. f x *\<^sub>R y) integrable_on A" | 
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 486 | using assms has_integral_scaleR_left unfolding integrable_on_def by blast | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 487 | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 488 | lemma has_integral_mult_left: | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 489 | fixes c :: "_ :: real_normed_algebra" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 490 | shows "(f has_integral y) S \<Longrightarrow> ((\<lambda>x. f x * c) has_integral (y * c)) S" | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 491 | using has_integral_linear[OF _ bounded_linear_mult_left] by (simp add: comp_def) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 492 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 493 | text\<open>The case analysis eliminates the condition @{term "f integrable_on S"} at the cost
 | 
| 62837 | 494 | of the type class constraint \<open>division_ring\<close>\<close> | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 495 | corollary integral_mult_left [simp]: | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 496 |   fixes c:: "'a::{real_normed_algebra,division_ring}"
 | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 497 | shows "integral S (\<lambda>x. f x * c) = integral S f * c" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 498 | proof (cases "f integrable_on S \<or> c = 0") | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 499 | case True then show ?thesis | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 500 | by (force intro: has_integral_mult_left) | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 501 | next | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 502 | case False then have "~ (\<lambda>x. f x * c) integrable_on S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 503 | using has_integral_mult_left [of "(\<lambda>x. f x * c)" _ S "inverse c"] | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 504 | by (auto simp add: mult.assoc) | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 505 | with False show ?thesis by (simp add: not_integrable_integral) | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 506 | qed | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 507 | |
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 508 | corollary integral_mult_right [simp]: | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 509 |   fixes c:: "'a::{real_normed_field}"
 | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 510 | shows "integral S (\<lambda>x. c * f x) = c * integral S f" | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 511 | by (simp add: mult.commute [of c]) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 512 | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 513 | corollary integral_divide [simp]: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 514 | fixes z :: "'a::real_normed_field" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 515 | shows "integral S (\<lambda>x. f x / z) = integral S (\<lambda>x. f x) / z" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 516 | using integral_mult_left [of S f "inverse z"] | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 517 | by (simp add: divide_inverse_commute) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 518 | |
| 60762 | 519 | lemma has_integral_mult_right: | 
| 520 | fixes c :: "'a :: real_normed_algebra" | |
| 521 | shows "(f has_integral y) i \<Longrightarrow> ((\<lambda>x. c * f x) has_integral (c * y)) i" | |
| 522 | using has_integral_linear[OF _ bounded_linear_mult_right] by (simp add: comp_def) | |
| 61165 | 523 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 524 | lemma has_integral_cmul: "(f has_integral k) S \<Longrightarrow> ((\<lambda>x. c *\<^sub>R f x) has_integral (c *\<^sub>R k)) S" | 
| 53410 | 525 | unfolding o_def[symmetric] | 
| 60396 | 526 | by (metis has_integral_linear bounded_linear_scaleR_right) | 
| 35172 | 527 | |
| 50104 | 528 | lemma has_integral_cmult_real: | 
| 529 | fixes c :: real | |
| 530 | assumes "c \<noteq> 0 \<Longrightarrow> (f has_integral x) A" | |
| 531 | shows "((\<lambda>x. c * f x) has_integral c * x) A" | |
| 53410 | 532 | proof (cases "c = 0") | 
| 533 | case True | |
| 534 | then show ?thesis by simp | |
| 535 | next | |
| 536 | case False | |
| 50104 | 537 | from has_integral_cmul[OF assms[OF this], of c] show ?thesis | 
| 538 | unfolding real_scaleR_def . | |
| 53410 | 539 | qed | 
| 540 | ||
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 541 | lemma has_integral_neg: "(f has_integral k) S \<Longrightarrow> ((\<lambda>x. -(f x)) has_integral -k) S" | 
| 60396 | 542 | by (drule_tac c="-1" in has_integral_cmul) auto | 
| 53410 | 543 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 544 | lemma has_integral_neg_iff: "((\<lambda>x. - f x) has_integral k) S \<longleftrightarrow> (f has_integral - k) S" | 
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 545 | using has_integral_neg[of f "- k"] has_integral_neg[of "\<lambda>x. - f x" k] by auto | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 546 | |
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 547 | lemma has_integral_add_cbox: | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 548 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 549 | assumes "(f has_integral k) (cbox a b)" "(g has_integral l) (cbox a b)" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 550 | shows "((\<lambda>x. f x + g x) has_integral (k + l)) (cbox a b)" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 551 | using assms | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 552 | unfolding has_integral_cbox | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 553 | by (simp add: split_beta' scaleR_add_right sum.distrib[abs_def] tendsto_add) | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 554 | |
| 53410 | 555 | lemma has_integral_add: | 
| 56188 | 556 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::real_normed_vector" | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 557 | assumes f: "(f has_integral k) S" and g: "(g has_integral l) S" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 558 | shows "((\<lambda>x. f x + g x) has_integral (k + l)) S" | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 559 | proof (cases "\<exists>a b. S = cbox a b") | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 560 | case True with has_integral_add_cbox assms show ?thesis | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 561 | by blast | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 562 | next | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 563 | let ?S = "\<lambda>f x. if x \<in> S then f x else 0" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 564 | case False | 
| 53410 | 565 | then show ?thesis | 
| 61166 
5976fe402824
renamed method "goals" to "goal_cases" to emphasize its meaning;
 wenzelm parents: 
61165diff
changeset | 566 | proof (subst has_integral_alt, clarsimp, goal_cases) | 
| 61165 | 567 | case (1 e) | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 568 | then have e2: "e/2 > 0" | 
| 53410 | 569 | by auto | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 570 | obtain Bf where "0 < Bf" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 571 | and Bf: "\<And>a b. ball 0 Bf \<subseteq> cbox a b \<Longrightarrow> | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 572 | \<exists>z. (?S f has_integral z) (cbox a b) \<and> norm (z - k) < e/2" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 573 | using has_integral_altD[OF f False e2] by blast | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 574 | obtain Bg where "0 < Bg" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 575 | and Bg: "\<And>a b. ball 0 Bg \<subseteq> (cbox a b) \<Longrightarrow> | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 576 | \<exists>z. (?S g has_integral z) (cbox a b) \<and> norm (z - l) < e/2" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 577 | using has_integral_altD[OF g False e2] by blast | 
| 53410 | 578 | show ?case | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 579 | proof (rule_tac x="max Bf Bg" in exI, clarsimp simp add: max.strict_coboundedI1 \<open>0 < Bf\<close>) | 
| 53410 | 580 | fix a b | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 581 | assume "ball 0 (max Bf Bg) \<subseteq> cbox a (b::'n)" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 582 | then have fs: "ball 0 Bf \<subseteq> cbox a (b::'n)" and gs: "ball 0 Bg \<subseteq> cbox a (b::'n)" | 
| 53410 | 583 | by auto | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 584 | obtain w where w: "(?S f has_integral w) (cbox a b)" "norm (w - k) < e/2" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 585 | using Bf[OF fs] by blast | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 586 | obtain z where z: "(?S g has_integral z) (cbox a b)" "norm (z - l) < e/2" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 587 | using Bg[OF gs] by blast | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 588 | have *: "\<And>x. (if x \<in> S then f x + g x else 0) = (?S f x) + (?S g x)" | 
| 53410 | 589 | by auto | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 590 | show "\<exists>z. (?S(\<lambda>x. f x + g x) has_integral z) (cbox a b) \<and> norm (z - (k + l)) < e" | 
| 53410 | 591 | apply (rule_tac x="w + z" in exI) | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 592 | apply (simp add: has_integral_add_cbox[OF w(1) z(1), unfolded *[symmetric]]) | 
| 53410 | 593 | using norm_triangle_ineq[of "w - k" "z - l"] w(2) z(2) | 
| 594 | apply (auto simp add: field_simps) | |
| 595 | done | |
| 596 | qed | |
| 597 | qed | |
| 598 | qed | |
| 35172 | 599 | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 600 | lemma has_integral_diff: | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 601 | "(f has_integral k) S \<Longrightarrow> (g has_integral l) S \<Longrightarrow> | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 602 | ((\<lambda>x. f x - g x) has_integral (k - l)) S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 603 | using has_integral_add[OF _ has_integral_neg, of f k S g l] | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 604 | by (auto simp: algebra_simps) | 
| 53410 | 605 | |
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 606 | lemma integral_0 [simp]: | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 607 | "integral S (\<lambda>x::'n::euclidean_space. 0::'m::real_normed_vector) = 0" | 
| 53410 | 608 | by (rule integral_unique has_integral_0)+ | 
| 609 | ||
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 610 | lemma integral_add: "f integrable_on S \<Longrightarrow> g integrable_on S \<Longrightarrow> | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 611 | integral S (\<lambda>x. f x + g x) = integral S f + integral S g" | 
| 60396 | 612 | by (rule integral_unique) (metis integrable_integral has_integral_add) | 
| 53410 | 613 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 614 | lemma integral_cmul [simp]: "integral S (\<lambda>x. c *\<^sub>R f x) = c *\<^sub>R integral S f" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 615 | proof (cases "f integrable_on S \<or> c = 0") | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 616 | case True with has_integral_cmul integrable_integral show ?thesis | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 617 | by fastforce | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 618 | next | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 619 | case False then have "~ (\<lambda>x. c *\<^sub>R f x) integrable_on S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 620 | using has_integral_cmul [of "(\<lambda>x. c *\<^sub>R f x)" _ S "inverse c"] by auto | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 621 | with False show ?thesis by (simp add: not_integrable_integral) | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 622 | qed | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 623 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 624 | lemma integral_neg [simp]: "integral S (\<lambda>x. - f x) = - integral S f" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 625 | proof (cases "f integrable_on S") | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 626 | case True then show ?thesis | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 627 | by (simp add: has_integral_neg integrable_integral integral_unique) | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 628 | next | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 629 | case False then have "~ (\<lambda>x. - f x) integrable_on S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 630 | using has_integral_neg [of "(\<lambda>x. - f x)" _ S ] by auto | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 631 | with False show ?thesis by (simp add: not_integrable_integral) | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 632 | qed | 
| 53410 | 633 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 634 | lemma integral_diff: "f integrable_on S \<Longrightarrow> g integrable_on S \<Longrightarrow> | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 635 | integral S (\<lambda>x. f x - g x) = integral S f - integral S g" | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 636 | by (rule integral_unique) (metis integrable_integral has_integral_diff) | 
| 35172 | 637 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 638 | lemma integrable_0: "(\<lambda>x. 0) integrable_on S" | 
| 35172 | 639 | unfolding integrable_on_def using has_integral_0 by auto | 
| 640 | ||
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 641 | lemma integrable_add: "f integrable_on S \<Longrightarrow> g integrable_on S \<Longrightarrow> (\<lambda>x. f x + g x) integrable_on S" | 
| 35172 | 642 | unfolding integrable_on_def by(auto intro: has_integral_add) | 
| 643 | ||
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 644 | lemma integrable_cmul: "f integrable_on S \<Longrightarrow> (\<lambda>x. c *\<^sub>R f(x)) integrable_on S" | 
| 35172 | 645 | unfolding integrable_on_def by(auto intro: has_integral_cmul) | 
| 646 | ||
| 50104 | 647 | lemma integrable_on_cmult_iff: | 
| 53410 | 648 | fixes c :: real | 
| 649 | assumes "c \<noteq> 0" | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 650 | shows "(\<lambda>x. c * f x) integrable_on S \<longleftrightarrow> f integrable_on S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 651 | using integrable_cmul[of "\<lambda>x. c * f x" S "1 / c"] integrable_cmul[of f S c] \<open>c \<noteq> 0\<close> | 
| 50104 | 652 | by auto | 
| 653 | ||
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 654 | lemma integrable_on_cmult_left: | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 655 | assumes "f integrable_on S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 656 | shows "(\<lambda>x. of_real c * f x) integrable_on S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 657 | using integrable_cmul[of f S "of_real c"] assms | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 658 | by (simp add: scaleR_conv_of_real) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 659 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 660 | lemma integrable_neg: "f integrable_on S \<Longrightarrow> (\<lambda>x. -f(x)) integrable_on S" | 
| 35172 | 661 | unfolding integrable_on_def by(auto intro: has_integral_neg) | 
| 662 | ||
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 663 | lemma integrable_diff: | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 664 | "f integrable_on S \<Longrightarrow> g integrable_on S \<Longrightarrow> (\<lambda>x. f x - g x) integrable_on S" | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 665 | unfolding integrable_on_def by(auto intro: has_integral_diff) | 
| 35172 | 666 | |
| 667 | lemma integrable_linear: | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 668 | "f integrable_on S \<Longrightarrow> bounded_linear h \<Longrightarrow> (h \<circ> f) integrable_on S" | 
| 35172 | 669 | unfolding integrable_on_def by(auto intro: has_integral_linear) | 
| 670 | ||
| 671 | lemma integral_linear: | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 672 | "f integrable_on S \<Longrightarrow> bounded_linear h \<Longrightarrow> integral S (h \<circ> f) = h (integral S f)" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 673 | apply (rule has_integral_unique [where i=S and f = "h \<circ> f"]) | 
| 60396 | 674 | apply (simp_all add: integrable_integral integrable_linear has_integral_linear ) | 
| 53410 | 675 | done | 
| 676 | ||
| 677 | lemma integral_component_eq[simp]: | |
| 56188 | 678 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 679 | assumes "f integrable_on S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 680 | shows "integral S (\<lambda>x. f x \<bullet> k) = integral S f \<bullet> k" | 
| 63938 | 681 | unfolding integral_linear[OF assms(1) bounded_linear_inner_left,unfolded o_def] .. | 
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 682 | |
| 64267 | 683 | lemma has_integral_sum: | 
| 66560 | 684 | assumes "finite T" | 
| 685 | and "\<And>a. a \<in> T \<Longrightarrow> ((f a) has_integral (i a)) S" | |
| 686 | shows "((\<lambda>x. sum (\<lambda>a. f a x) T) has_integral (sum i T)) S" | |
| 687 | using assms(1) subset_refl[of T] | |
| 53410 | 688 | proof (induct rule: finite_subset_induct) | 
| 689 | case empty | |
| 690 | then show ?case by auto | |
| 691 | next | |
| 692 | case (insert x F) | |
| 60396 | 693 | with assms show ?case | 
| 694 | by (simp add: has_integral_add) | |
| 695 | qed | |
| 696 | ||
| 64267 | 697 | lemma integral_sum: | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 698 | "\<lbrakk>finite I; \<And>a. a \<in> I \<Longrightarrow> f a integrable_on S\<rbrakk> \<Longrightarrow> | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 699 | integral S (\<lambda>x. \<Sum>a\<in>I. f a x) = (\<Sum>a\<in>I. integral S (f a))" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 700 | by (simp add: has_integral_sum integrable_integral integral_unique) | 
| 64267 | 701 | |
| 702 | lemma integrable_sum: | |
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 703 | "\<lbrakk>finite I; \<And>a. a \<in> I \<Longrightarrow> f a integrable_on S\<rbrakk> \<Longrightarrow> (\<lambda>x. \<Sum>a\<in>I. f a x) integrable_on S" | 
| 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 704 | unfolding integrable_on_def using has_integral_sum[of I] by metis | 
| 35172 | 705 | |
| 706 | lemma has_integral_eq: | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 707 | assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x" | 
| 53410 | 708 | and "(f has_integral k) s" | 
| 709 | shows "(g has_integral k) s" | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 710 | using has_integral_diff[OF assms(2), of "\<lambda>x. f x - g x" 0] | 
| 53410 | 711 | using has_integral_is_0[of s "\<lambda>x. f x - g x"] | 
| 712 | using assms(1) | |
| 713 | by auto | |
| 714 | ||
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 715 | lemma integrable_eq: "\<lbrakk>f integrable_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g integrable_on s" | 
| 53410 | 716 | unfolding integrable_on_def | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 717 | using has_integral_eq[of s f g] has_integral_eq by blast | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 718 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 719 | lemma has_integral_cong: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 720 | assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 721 | shows "(f has_integral i) s = (g has_integral i) s" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 722 | using has_integral_eq[of s f g] has_integral_eq[of s g f] assms | 
| 53410 | 723 | by auto | 
| 724 | ||
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 725 | lemma integral_cong: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 726 | assumes "\<And>x. x \<in> s \<Longrightarrow> f x = g x" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 727 | shows "integral s f = integral s g" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 728 | unfolding integral_def | 
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 729 | by (metis (full_types, hide_lams) assms has_integral_cong integrable_eq) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 730 | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 731 | lemma integrable_on_cmult_left_iff [simp]: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 732 | assumes "c \<noteq> 0" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 733 | shows "(\<lambda>x. of_real c * f x) integrable_on s \<longleftrightarrow> f integrable_on s" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 734 | (is "?lhs = ?rhs") | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 735 | proof | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 736 | assume ?lhs | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 737 | then have "(\<lambda>x. of_real (1 / c) * (of_real c * f x)) integrable_on s" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 738 | using integrable_cmul[of "\<lambda>x. of_real c * f x" s "1 / of_real c"] | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 739 | by (simp add: scaleR_conv_of_real) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 740 | then have "(\<lambda>x. (of_real (1 / c) * of_real c * f x)) integrable_on s" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 741 | by (simp add: algebra_simps) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 742 | with \<open>c \<noteq> 0\<close> show ?rhs | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 743 | by (metis (no_types, lifting) integrable_eq mult.left_neutral nonzero_divide_eq_eq of_real_1 of_real_mult) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 744 | qed (blast intro: integrable_on_cmult_left) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 745 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 746 | lemma integrable_on_cmult_right: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 747 |   fixes f :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 748 | assumes "f integrable_on s" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 749 | shows "(\<lambda>x. f x * of_real c) integrable_on s" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 750 | using integrable_on_cmult_left [OF assms] by (simp add: mult.commute) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 751 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 752 | lemma integrable_on_cmult_right_iff [simp]: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 753 |   fixes f :: "_ \<Rightarrow> 'b :: {comm_ring,real_algebra_1,real_normed_vector}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 754 | assumes "c \<noteq> 0" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 755 | shows "(\<lambda>x. f x * of_real c) integrable_on s \<longleftrightarrow> f integrable_on s" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 756 | using integrable_on_cmult_left_iff [OF assms] by (simp add: mult.commute) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 757 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 758 | lemma integrable_on_cdivide: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 759 | fixes f :: "_ \<Rightarrow> 'b :: real_normed_field" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 760 | assumes "f integrable_on s" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 761 | shows "(\<lambda>x. f x / of_real c) integrable_on s" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 762 | by (simp add: integrable_on_cmult_right divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 763 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 764 | lemma integrable_on_cdivide_iff [simp]: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 765 | fixes f :: "_ \<Rightarrow> 'b :: real_normed_field" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 766 | assumes "c \<noteq> 0" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 767 | shows "(\<lambda>x. f x / of_real c) integrable_on s \<longleftrightarrow> f integrable_on s" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 768 | by (simp add: divide_inverse assms of_real_inverse [symmetric] del: of_real_inverse) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 769 | |
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 770 | lemma has_integral_null [intro]: "content(cbox a b) = 0 \<Longrightarrow> (f has_integral 0) (cbox a b)" | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 771 | unfolding has_integral_cbox | 
| 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 772 | using eventually_division_filter_tagged_division[of "cbox a b"] | 
| 64267 | 773 | by (subst tendsto_cong[where g="\<lambda>_. 0"]) (auto elim: eventually_mono intro: sum_content_null) | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 774 | |
| 66402 | 775 | lemma has_integral_null_real [intro]: "content {a..b::real} = 0 \<Longrightarrow> (f has_integral 0) {a..b}"
 | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 776 | by (metis box_real(2) has_integral_null) | 
| 56188 | 777 | |
| 778 | lemma has_integral_null_eq[simp]: "content (cbox a b) = 0 \<Longrightarrow> (f has_integral i) (cbox a b) \<longleftrightarrow> i = 0" | |
| 60396 | 779 | by (auto simp add: has_integral_null dest!: integral_unique) | 
| 53410 | 780 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 781 | lemma integral_null [simp]: "content (cbox a b) = 0 \<Longrightarrow> integral (cbox a b) f = 0" | 
| 60396 | 782 | by (metis has_integral_null integral_unique) | 
| 53410 | 783 | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 784 | lemma integrable_on_null [intro]: "content (cbox a b) = 0 \<Longrightarrow> f integrable_on (cbox a b)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 785 | by (simp add: has_integral_integrable) | 
| 53410 | 786 | |
| 787 | lemma has_integral_empty[intro]: "(f has_integral 0) {}"
 | |
| 66519 | 788 | by (meson ex_in_conv has_integral_is_0) | 
| 53410 | 789 | |
| 790 | lemma has_integral_empty_eq[simp]: "(f has_integral i) {} \<longleftrightarrow> i = 0"
 | |
| 60396 | 791 | by (auto simp add: has_integral_empty has_integral_unique) | 
| 53410 | 792 | |
| 793 | lemma integrable_on_empty[intro]: "f integrable_on {}"
 | |
| 794 | unfolding integrable_on_def by auto | |
| 795 | ||
| 796 | lemma integral_empty[simp]: "integral {} f = 0"
 | |
| 797 | by (rule integral_unique) (rule has_integral_empty) | |
| 798 | ||
| 799 | lemma has_integral_refl[intro]: | |
| 56188 | 800 | fixes a :: "'a::euclidean_space" | 
| 801 | shows "(f has_integral 0) (cbox a a)" | |
| 53410 | 802 |     and "(f has_integral 0) {a}"
 | 
| 803 | proof - | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 804 | show "(f has_integral 0) (cbox a a)" | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 805 | by (rule has_integral_null) simp | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 806 |   then show "(f has_integral 0) {a}"
 | 
| 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 807 | by simp | 
| 56188 | 808 | qed | 
| 809 | ||
| 810 | lemma integrable_on_refl[intro]: "f integrable_on cbox a a" | |
| 53410 | 811 | unfolding integrable_on_def by auto | 
| 812 | ||
| 60762 | 813 | lemma integral_refl [simp]: "integral (cbox a a) f = 0" | 
| 53410 | 814 | by (rule integral_unique) auto | 
| 815 | ||
| 60762 | 816 | lemma integral_singleton [simp]: "integral {a} f = 0"
 | 
| 817 | by auto | |
| 818 | ||
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 819 | lemma integral_blinfun_apply: | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 820 | assumes "f integrable_on s" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 821 | shows "integral s (\<lambda>x. blinfun_apply h (f x)) = blinfun_apply h (integral s f)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 822 | by (subst integral_linear[symmetric, OF assms blinfun.bounded_linear_right]) (simp add: o_def) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 823 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 824 | lemma blinfun_apply_integral: | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 825 | assumes "f integrable_on s" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 826 | shows "blinfun_apply (integral s f) x = integral s (\<lambda>y. blinfun_apply (f y) x)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 827 | by (metis (no_types, lifting) assms blinfun.prod_left.rep_eq integral_blinfun_apply integral_cong) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 828 | |
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 829 | lemma has_integral_componentwise_iff: | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 830 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 831 | shows "(f has_integral y) A \<longleftrightarrow> (\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 832 | proof safe | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 833 | fix b :: 'b assume "(f has_integral y) A" | 
| 63938 | 834 | from has_integral_linear[OF this(1) bounded_linear_inner_left, of b] | 
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 835 | show "((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A" by (simp add: o_def) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 836 | next | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 837 | assume "(\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 838 | hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral ((y \<bullet> b) *\<^sub>R b)) A" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 839 | by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 840 | hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. (y \<bullet> b) *\<^sub>R b)) A" | 
| 64267 | 841 | by (intro has_integral_sum) (simp_all add: o_def) | 
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 842 | thus "(f has_integral y) A" by (simp add: euclidean_representation) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 843 | qed | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 844 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 845 | lemma has_integral_componentwise: | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 846 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 847 | shows "(\<And>b. b \<in> Basis \<Longrightarrow> ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A) \<Longrightarrow> (f has_integral y) A" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 848 | by (subst has_integral_componentwise_iff) blast | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 849 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 850 | lemma integrable_componentwise_iff: | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 851 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 852 | shows "f integrable_on A \<longleftrightarrow> (\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 853 | proof | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 854 | assume "f integrable_on A" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 855 | then obtain y where "(f has_integral y) A" by (auto simp: integrable_on_def) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 856 | hence "(\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral (y \<bullet> b)) A)" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 857 | by (subst (asm) has_integral_componentwise_iff) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 858 | thus "(\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)" by (auto simp: integrable_on_def) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 859 | next | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 860 | assume "(\<forall>b\<in>Basis. (\<lambda>x. f x \<bullet> b) integrable_on A)" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 861 | then obtain y where "\<forall>b\<in>Basis. ((\<lambda>x. f x \<bullet> b) has_integral y b) A" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 862 | unfolding integrable_on_def by (subst (asm) bchoice_iff) blast | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 863 | hence "\<forall>b\<in>Basis. (((\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. f x \<bullet> b)) has_integral (y b *\<^sub>R b)) A" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 864 | by (intro ballI has_integral_linear) (simp_all add: bounded_linear_scaleR_left) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 865 | hence "((\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b) has_integral (\<Sum>b\<in>Basis. y b *\<^sub>R b)) A" | 
| 64267 | 866 | by (intro has_integral_sum) (simp_all add: o_def) | 
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 867 | thus "f integrable_on A" by (auto simp: integrable_on_def o_def euclidean_representation) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 868 | qed | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 869 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 870 | lemma integrable_componentwise: | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 871 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 872 | shows "(\<And>b. b \<in> Basis \<Longrightarrow> (\<lambda>x. f x \<bullet> b) integrable_on A) \<Longrightarrow> f integrable_on A" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 873 | by (subst integrable_componentwise_iff) blast | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 874 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 875 | lemma integral_componentwise: | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 876 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: euclidean_space" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 877 | assumes "f integrable_on A" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 878 | shows "integral A f = (\<Sum>b\<in>Basis. integral A (\<lambda>x. (f x \<bullet> b) *\<^sub>R b))" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 879 | proof - | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 880 | from assms have integrable: "\<forall>b\<in>Basis. (\<lambda>x. x *\<^sub>R b) \<circ> (\<lambda>x. (f x \<bullet> b)) integrable_on A" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 881 | by (subst (asm) integrable_componentwise_iff, intro integrable_linear ballI) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 882 | (simp_all add: bounded_linear_scaleR_left) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 883 | have "integral A f = integral A (\<lambda>x. \<Sum>b\<in>Basis. (f x \<bullet> b) *\<^sub>R b)" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 884 | by (simp add: euclidean_representation) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 885 | also from integrable have "\<dots> = (\<Sum>a\<in>Basis. integral A (\<lambda>x. (f x \<bullet> a) *\<^sub>R a))" | 
| 64267 | 886 | by (subst integral_sum) (simp_all add: o_def) | 
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 887 | finally show ?thesis . | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 888 | qed | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 889 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 890 | lemma integrable_component: | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 891 | "f integrable_on A \<Longrightarrow> (\<lambda>x. f x \<bullet> (y :: 'b :: euclidean_space)) integrable_on A" | 
| 63938 | 892 | by (drule integrable_linear[OF _ bounded_linear_inner_left[of y]]) (simp add: o_def) | 
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 893 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 894 | |
| 35172 | 895 | |
| 60420 | 896 | subsection \<open>Cauchy-type criterion for integrability.\<close> | 
| 35172 | 897 | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 898 | proposition integrable_Cauchy: | 
| 56188 | 899 |   fixes f :: "'n::euclidean_space \<Rightarrow> 'a::{real_normed_vector,complete_space}"
 | 
| 900 | shows "f integrable_on cbox a b \<longleftrightarrow> | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 901 | (\<forall>e>0. \<exists>\<gamma>. gauge \<gamma> \<and> | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 902 | (\<forall>\<D>1 \<D>2. \<D>1 tagged_division_of (cbox a b) \<and> \<gamma> fine \<D>1 \<and> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 903 | \<D>2 tagged_division_of (cbox a b) \<and> \<gamma> fine \<D>2 \<longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 904 | norm ((\<Sum>(x,K)\<in>\<D>1. content K *\<^sub>R f x) - (\<Sum>(x,K)\<in>\<D>2. content K *\<^sub>R f x)) < e))" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 905 | (is "?l = (\<forall>e>0. \<exists>\<gamma>. ?P e \<gamma>)") | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 906 | proof (intro iffI allI impI) | 
| 53442 | 907 | assume ?l | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 908 | then obtain y | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 909 | where y: "\<And>e. e > 0 \<Longrightarrow> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 910 | \<exists>\<gamma>. gauge \<gamma> \<and> | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 911 | (\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 912 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - y) < e)" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 913 | by (auto simp: integrable_on_def has_integral) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 914 | show "\<exists>\<gamma>. ?P e \<gamma>" if "e > 0" for e | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 915 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 916 | have "e/2 > 0" using that by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 917 | with y obtain \<gamma> where "gauge \<gamma>" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 918 | and \<gamma>: "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<Longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 919 | norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f x) - y) < e/2" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 920 | by meson | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 921 | show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 922 | apply (rule_tac x=\<gamma> in exI, clarsimp simp: \<open>gauge \<gamma>\<close>) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 923 | by (blast intro!: \<gamma> dist_triangle_half_l[where y=y,unfolded dist_norm]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 924 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 925 | next | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 926 | assume "\<forall>e>0. \<exists>\<gamma>. ?P e \<gamma>" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 927 | then have "\<forall>n::nat. \<exists>\<gamma>. ?P (1 / (n + 1)) \<gamma>" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 928 | by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 929 | then obtain \<gamma> :: "nat \<Rightarrow> 'n \<Rightarrow> 'n set" where \<gamma>: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 930 | "\<And>m. gauge (\<gamma> m)" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 931 | "\<And>m \<D>1 \<D>2. \<lbrakk>\<D>1 tagged_division_of cbox a b; | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 932 | \<gamma> m fine \<D>1; \<D>2 tagged_division_of cbox a b; \<gamma> m fine \<D>2\<rbrakk> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 933 | \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>1. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> \<D>2. content K *\<^sub>R f x)) | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 934 | < 1 / (m + 1)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 935 | by metis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 936 |   have "\<And>n. gauge (\<lambda>x. \<Inter>{\<gamma> i x |i. i \<in> {0..n}})"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 937 | apply (rule gauge_Inter) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 938 | using \<gamma> by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 939 |   then have "\<forall>n. \<exists>p. p tagged_division_of (cbox a b) \<and> (\<lambda>x. \<Inter>{\<gamma> i x |i. i \<in> {0..n}}) fine p"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 940 | by (meson fine_division_exists) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 941 | then obtain p where p: "\<And>z. p z tagged_division_of cbox a b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 942 |                          "\<And>z. (\<lambda>x. \<Inter>{\<gamma> i x |i. i \<in> {0..z}}) fine p z"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 943 | by meson | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 944 | have dp: "\<And>i n. i\<le>n \<Longrightarrow> \<gamma> i fine p n" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 945 | using p unfolding fine_Inter | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 946 | using atLeastAtMost_iff by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 947 | have "Cauchy (\<lambda>n. sum (\<lambda>(x,K). content K *\<^sub>R (f x)) (p n))" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 948 | proof (rule CauchyI) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 949 | fix e::real | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 950 | assume "0 < e" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 951 | then obtain N where "N \<noteq> 0" and N: "inverse (real N) < e" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 952 | using real_arch_inverse[of e] by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 953 | show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. norm ((\<Sum>(x,K) \<in> p m. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> p n. content K *\<^sub>R f x)) < e" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 954 | proof (intro exI allI impI) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 955 | fix m n | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 956 | assume mn: "N \<le> m" "N \<le> n" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 957 | have "norm ((\<Sum>(x,K) \<in> p m. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> p n. content K *\<^sub>R f x)) < 1 / (real N + 1)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 958 | by (simp add: p(1) dp mn \<gamma>) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 959 | also have "... < e" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 960 | using N \<open>N \<noteq> 0\<close> \<open>0 < e\<close> by (auto simp: field_simps) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 961 | finally show "norm ((\<Sum>(x,K) \<in> p m. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> p n. content K *\<^sub>R f x)) < e" . | 
| 53442 | 962 | qed | 
| 963 | qed | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 964 | then obtain y where y: "\<exists>no. \<forall>n\<ge>no. norm ((\<Sum>(x,K) \<in> p n. content K *\<^sub>R f x) - y) < r" if "r > 0" for r | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 965 | by (auto simp: convergent_eq_Cauchy[symmetric] dest: LIMSEQ_D) | 
| 53442 | 966 | show ?l | 
| 967 | unfolding integrable_on_def has_integral | |
| 60425 | 968 | proof (rule_tac x=y in exI, clarify) | 
| 53442 | 969 | fix e :: real | 
| 970 | assume "e>0" | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 971 | then have e2: "e/2 > 0" by auto | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 972 | then obtain N1::nat where N1: "N1 \<noteq> 0" "inverse (real N1) < e/2" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 973 | using real_arch_inverse by blast | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 974 | obtain N2::nat where N2: "\<And>n. n \<ge> N2 \<Longrightarrow> norm ((\<Sum>(x,K) \<in> p n. content K *\<^sub>R f x) - y) < e/2" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 975 | using y[OF e2] by metis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 976 | show "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 977 | (\<forall>\<D>. \<D> tagged_division_of (cbox a b) \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 978 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - y) < e)" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 979 | proof (intro exI conjI allI impI) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 980 | show "gauge (\<gamma> (N1+N2))" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 981 | using \<gamma> by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 982 | show "norm ((\<Sum>(x,K) \<in> q. content K *\<^sub>R f x) - y) < e" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 983 | if "q tagged_division_of cbox a b \<and> \<gamma> (N1+N2) fine q" for q | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 984 | proof (rule norm_triangle_half_r) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 985 | have "norm ((\<Sum>(x,K) \<in> p (N1+N2). content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> q. content K *\<^sub>R f x)) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 986 | < 1 / (real (N1+N2) + 1)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 987 | by (rule \<gamma>; simp add: dp p that) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 988 | also have "... < e/2" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 989 | using N1 \<open>0 < e\<close> by (auto simp: field_simps intro: less_le_trans) | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 990 | finally show "norm ((\<Sum>(x,K) \<in> p (N1+N2). content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> q. content K *\<^sub>R f x)) < e/2" . | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 991 | show "norm ((\<Sum>(x,K) \<in> p (N1+N2). content K *\<^sub>R f x) - y) < e/2" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 992 | using N2 le_add_same_cancel2 by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 993 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 994 | qed | 
| 53442 | 995 | qed | 
| 996 | qed | |
| 997 | ||
| 35172 | 998 | |
| 60420 | 999 | subsection \<open>Additivity of integral on abutting intervals.\<close> | 
| 35172 | 1000 | |
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 1001 | lemma tagged_division_split_left_inj_content: | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1002 | assumes \<D>: "\<D> tagged_division_of S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1003 |     and "(x1, K1) \<in> \<D>" "(x2, K2) \<in> \<D>" "K1 \<noteq> K2" "K1 \<inter> {x. x\<bullet>k \<le> c} = K2 \<inter> {x. x\<bullet>k \<le> c}" "k \<in> Basis"
 | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1004 |   shows "content (K1 \<inter> {x. x\<bullet>k \<le> c}) = 0"
 | 
| 53443 | 1005 | proof - | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1006 | from tagged_division_ofD(4)[OF \<D> \<open>(x1, K1) \<in> \<D>\<close>] obtain a b where K1: "K1 = cbox a b" | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 1007 | by auto | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1008 |   then have "interior (K1 \<inter> {x. x \<bullet> k \<le> c}) = {}"
 | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 1009 | by (metis tagged_division_split_left_inj assms) | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1010 | then show ?thesis | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1011 | unfolding K1 interval_split[OF \<open>k \<in> Basis\<close>] by (auto simp: content_eq_0_interior) | 
| 53443 | 1012 | qed | 
| 1013 | ||
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 1014 | lemma tagged_division_split_right_inj_content: | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1015 | assumes \<D>: "\<D> tagged_division_of S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1016 |     and "(x1, K1) \<in> \<D>" "(x2, K2) \<in> \<D>" "K1 \<noteq> K2" "K1 \<inter> {x. x\<bullet>k \<ge> c} = K2 \<inter> {x. x\<bullet>k \<ge> c}" "k \<in> Basis"
 | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1017 |   shows "content (K1 \<inter> {x. x\<bullet>k \<ge> c}) = 0"
 | 
| 53443 | 1018 | proof - | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1019 | from tagged_division_ofD(4)[OF \<D> \<open>(x1, K1) \<in> \<D>\<close>] obtain a b where K1: "K1 = cbox a b" | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 1020 | by auto | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1021 |   then have "interior (K1 \<inter> {x. c \<le> x \<bullet> k}) = {}"
 | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 1022 | by (metis tagged_division_split_right_inj assms) | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1023 | then show ?thesis | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1024 | unfolding K1 interval_split[OF \<open>k \<in> Basis\<close>] | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1025 | by (auto simp: content_eq_0_interior) | 
| 53443 | 1026 | qed | 
| 35172 | 1027 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1028 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1029 | proposition has_integral_split: | 
| 56188 | 1030 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 60435 | 1031 |   assumes fi: "(f has_integral i) (cbox a b \<inter> {x. x\<bullet>k \<le> c})"
 | 
| 1032 |       and fj: "(f has_integral j) (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"
 | |
| 1033 | and k: "k \<in> Basis" | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1034 | shows "(f has_integral (i + j)) (cbox a b)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1035 | unfolding has_integral | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1036 | proof clarify | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1037 | fix e::real | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1038 | assume "0 < e" | 
| 53468 | 1039 | then have e: "e/2 > 0" | 
| 1040 | by auto | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1041 | obtain \<gamma>1 where \<gamma>1: "gauge \<gamma>1" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1042 | and \<gamma>1norm: | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 1043 |         "\<And>\<D>. \<lbrakk>\<D> tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<le> c}; \<gamma>1 fine \<D>\<rbrakk>
 | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 1044 | \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - i) < e/2" | 
| 60435 | 1045 | apply (rule has_integralD[OF fi[unfolded interval_split[OF k]] e]) | 
| 1046 | apply (simp add: interval_split[symmetric] k) | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1047 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1048 | obtain \<gamma>2 where \<gamma>2: "gauge \<gamma>2" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1049 | and \<gamma>2norm: | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 1050 |         "\<And>\<D>. \<lbrakk>\<D> tagged_division_of cbox a b \<inter> {x. c \<le> x \<bullet> k}; \<gamma>2 fine \<D>\<rbrakk>
 | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 1051 | \<Longrightarrow> norm ((\<Sum>(x, k) \<in> \<D>. content k *\<^sub>R f x) - j) < e/2" | 
| 60435 | 1052 | apply (rule has_integralD[OF fj[unfolded interval_split[OF k]] e]) | 
| 1053 | apply (simp add: interval_split[symmetric] k) | |
| 1054 | done | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1055 | let ?\<gamma> = "\<lambda>x. if x\<bullet>k = c then (\<gamma>1 x \<inter> \<gamma>2 x) else ball x \<bar>x\<bullet>k - c\<bar> \<inter> \<gamma>1 x \<inter> \<gamma>2 x" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1056 | have "gauge ?\<gamma>" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1057 | using \<gamma>1 \<gamma>2 unfolding gauge_def by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1058 | then show "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 1059 | (\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 1060 | norm ((\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R f x) - (i + j)) < e)" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1061 | proof (rule_tac x="?\<gamma>" in exI, safe) | 
| 53468 | 1062 | fix p | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1063 | assume p: "p tagged_division_of (cbox a b)" and "?\<gamma> fine p" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1064 |     have ab_eqp: "cbox a b = \<Union>{K. \<exists>x. (x, K) \<in> p}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1065 | using p by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1066 |     have xk_le_c: "x\<bullet>k \<le> c" if as: "(x,K) \<in> p" and K: "K \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}" for x K
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1067 | proof (rule ccontr) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1068 | assume **: "\<not> x \<bullet> k \<le> c" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1069 | then have "K \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1070 | using \<open>?\<gamma> fine p\<close> as by (fastforce simp: not_le algebra_simps) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1071 | with K obtain y where y: "y \<in> ball x \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<le> c" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1072 | by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1073 | then have "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1074 | using Basis_le_norm[OF k, of "x - y"] | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1075 | by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1076 | with y show False | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1077 | using ** by (auto simp add: field_simps) | 
| 60435 | 1078 | qed | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1079 |     have xk_ge_c: "x\<bullet>k \<ge> c" if as: "(x,K) \<in> p" and K: "K \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}" for x K
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1080 | proof (rule ccontr) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1081 | assume **: "\<not> x \<bullet> k \<ge> c" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1082 | then have "K \<subseteq> ball x \<bar>x \<bullet> k - c\<bar>" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1083 | using \<open>?\<gamma> fine p\<close> as by (fastforce simp: not_le algebra_simps) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1084 | with K obtain y where y: "y \<in> ball x \<bar>x \<bullet> k - c\<bar>" "y\<bullet>k \<ge> c" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1085 | by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1086 | then have "\<bar>x \<bullet> k - y \<bullet> k\<bar> < \<bar>x \<bullet> k - c\<bar>" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1087 | using Basis_le_norm[OF k, of "x - y"] | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1088 | by (auto simp add: dist_norm inner_diff_left intro: le_less_trans) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1089 | with y show False | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1090 | using ** by (auto simp add: field_simps) | 
| 53468 | 1091 | qed | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1092 |     have fin_finite: "finite {(x,f K) | x K. (x,K) \<in> s \<and> P x K}"
 | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 1093 | if "finite s" for s and f :: "'a set \<Rightarrow> 'a set" and P :: "'a \<Rightarrow> 'a set \<Rightarrow> bool" | 
| 53468 | 1094 | proof - | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1095 | from that have "finite ((\<lambda>(x,K). (x, f K)) ` s)" | 
| 60425 | 1096 | by auto | 
| 61165 | 1097 | then show ?thesis | 
| 60425 | 1098 | by (rule rev_finite_subset) auto | 
| 53468 | 1099 | qed | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1100 |     { fix \<G> :: "'a set \<Rightarrow> 'a set"
 | 
| 53468 | 1101 | fix i :: "'a \<times> 'a set" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1102 |       assume "i \<in> (\<lambda>(x, k). (x, \<G> k)) ` p - {(x, \<G> k) |x k. (x, k) \<in> p \<and> \<G> k \<noteq> {}}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1103 | then obtain x K where xk: "i = (x, \<G> K)" "(x,K) \<in> p" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1104 |                                  "(x, \<G> K) \<notin> {(x, \<G> K) |x K. (x,K) \<in> p \<and> \<G> K \<noteq> {}}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1105 | by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1106 | have "content (\<G> K) = 0" | 
| 53468 | 1107 | using xk using content_empty by auto | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1108 | then have "(\<lambda>(x,K). content K *\<^sub>R f x) i = 0" | 
| 53468 | 1109 | unfolding xk split_conv by auto | 
| 60435 | 1110 | } note [simp] = this | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1111 | have "finite p" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1112 | using p by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1113 |     let ?M1 = "{(x, K \<inter> {x. x\<bullet>k \<le> c}) |x K. (x,K) \<in> p \<and> K \<inter> {x. x\<bullet>k \<le> c} \<noteq> {}}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1114 | have \<gamma>1_fine: "\<gamma>1 fine ?M1" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1115 | using \<open>?\<gamma> fine p\<close> by (fastforce simp: fine_def split: if_split_asm) | 
| 53468 | 1116 | have "norm ((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) < e/2" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1117 | proof (rule \<gamma>1norm [OF tagged_division_ofI \<gamma>1_fine]) | 
| 60435 | 1118 | show "finite ?M1" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1119 | by (rule fin_finite) (use p in blast) | 
| 56188 | 1120 |       show "\<Union>{k. \<exists>x. (x, k) \<in> ?M1} = cbox a b \<inter> {x. x\<bullet>k \<le> c}"
 | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1121 | by (auto simp: ab_eqp) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1122 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1123 | fix x L | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1124 | assume xL: "(x, L) \<in> ?M1" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1125 |       then obtain x' L' where xL': "x = x'" "L = L' \<inter> {x. x \<bullet> k \<le> c}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1126 |                                    "(x', L') \<in> p" "L' \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1127 | by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1128 | then obtain a' b' where ab': "L' = cbox a' b'" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1129 | using p by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1130 |       show "x \<in> L" "L \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<le> c}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1131 | using p xk_le_c xL' by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1132 | show "\<exists>a b. L = cbox a b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1133 | using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1134 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1135 | fix y R | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1136 | assume yR: "(y, R) \<in> ?M1" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1137 |       then obtain y' R' where yR': "y = y'" "R = R' \<inter> {x. x \<bullet> k \<le> c}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1138 |                                    "(y', R') \<in> p" "R' \<inter> {x. x \<bullet> k \<le> c} \<noteq> {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1139 | by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1140 | assume as: "(x, L) \<noteq> (y, R)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1141 |       show "interior L \<inter> interior R = {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1142 | proof (cases "L' = R' \<longrightarrow> x' = y'") | 
| 53468 | 1143 | case False | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1144 |         have "interior R' = {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1145 | by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) | 
| 53468 | 1146 | then show ?thesis | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1147 | using yR' by simp | 
| 53468 | 1148 | next | 
| 1149 | case True | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1150 | then have "L' \<noteq> R'" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1151 | using as unfolding xL' yR' by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1152 |         have "interior L' \<inter> interior R' = {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1153 | by (metis (no_types) Pair_inject \<open>L' \<noteq> R'\<close> p tagged_division_ofD(5) xL'(3) yR'(3)) | 
| 53468 | 1154 | then show ?thesis | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1155 | using xL'(2) yR'(2) by auto | 
| 35172 | 1156 | qed | 
| 1157 | qed | |
| 53468 | 1158 | moreover | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1159 |     let ?M2 = "{(x,K \<inter> {x. x\<bullet>k \<ge> c}) |x K. (x,K) \<in> p \<and> K \<inter> {x. x\<bullet>k \<ge> c} \<noteq> {}}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1160 | have \<gamma>2_fine: "\<gamma>2 fine ?M2" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1161 | using \<open>?\<gamma> fine p\<close> by (fastforce simp: fine_def split: if_split_asm) | 
| 53468 | 1162 | have "norm ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) < e/2" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1163 | proof (rule \<gamma>2norm [OF tagged_division_ofI \<gamma>2_fine]) | 
| 60435 | 1164 | show "finite ?M2" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1165 | by (rule fin_finite) (use p in blast) | 
| 56188 | 1166 |       show "\<Union>{k. \<exists>x. (x, k) \<in> ?M2} = cbox a b \<inter> {x. x\<bullet>k \<ge> c}"
 | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1167 | by (auto simp: ab_eqp) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1168 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1169 | fix x L | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1170 | assume xL: "(x, L) \<in> ?M2" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1171 |       then obtain x' L' where xL': "x = x'" "L = L' \<inter> {x. x \<bullet> k \<ge> c}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1172 |                                    "(x', L') \<in> p" "L' \<inter> {x. x \<bullet> k \<ge> c} \<noteq> {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1173 | by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1174 | then obtain a' b' where ab': "L' = cbox a' b'" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1175 | using p by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1176 |       show "x \<in> L" "L \<subseteq> cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1177 | using p xk_ge_c xL' by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1178 | show "\<exists>a b. L = cbox a b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1179 | using p xL' ab' by (auto simp add: interval_split[OF k,where c=c]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1180 | |
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1181 | fix y R | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1182 | assume yR: "(y, R) \<in> ?M2" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1183 |       then obtain y' R' where yR': "y = y'" "R = R' \<inter> {x. x \<bullet> k \<ge> c}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1184 |                                    "(y', R') \<in> p" "R' \<inter> {x. x \<bullet> k \<ge> c} \<noteq> {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1185 | by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1186 | assume as: "(x, L) \<noteq> (y, R)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1187 |       show "interior L \<inter> interior R = {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1188 | proof (cases "L' = R' \<longrightarrow> x' = y'") | 
| 53468 | 1189 | case False | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1190 |         have "interior R' = {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1191 | by (metis (no_types) False Pair_inject inf.idem tagged_division_ofD(5) [OF p] xL'(3) yR'(3)) | 
| 53468 | 1192 | then show ?thesis | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1193 | using yR' by simp | 
| 53468 | 1194 | next | 
| 1195 | case True | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1196 | then have "L' \<noteq> R'" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1197 | using as unfolding xL' yR' by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1198 |         have "interior L' \<inter> interior R' = {}"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1199 | by (metis (no_types) Pair_inject \<open>L' \<noteq> R'\<close> p tagged_division_ofD(5) xL'(3) yR'(3)) | 
| 53468 | 1200 | then show ?thesis | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1201 | using xL'(2) yR'(2) by auto | 
| 53468 | 1202 | qed | 
| 1203 | qed | |
| 1204 | ultimately | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1205 | have "norm (((\<Sum>(x,K) \<in> ?M1. content K *\<^sub>R f x) - i) + ((\<Sum>(x,K) \<in> ?M2. content K *\<^sub>R f x) - j)) < e/2 + e/2" | 
| 60425 | 1206 | using norm_add_less by blast | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1207 | moreover have "((\<Sum>(x,K) \<in> ?M1. content K *\<^sub>R f x) - i) + | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1208 | ((\<Sum>(x,K) \<in> ?M2. content K *\<^sub>R f x) - j) = | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1209 | (\<Sum>(x, ka)\<in>p. content ka *\<^sub>R f x) - (i + j)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1210 | proof - | 
| 60435 | 1211 | have eq0: "\<And>x y. x = (0::real) \<Longrightarrow> x *\<^sub>R (y::'b) = 0" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1212 | by auto | 
| 60435 | 1213 | have cont_eq: "\<And>g. (\<lambda>(x,l). content l *\<^sub>R f x) \<circ> (\<lambda>(x,l). (x,g l)) = (\<lambda>(x,l). content (g l) *\<^sub>R f x)" | 
| 1214 | by auto | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1215 | have *: "\<And>\<G> :: 'a set \<Rightarrow> 'a set. | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1216 |                   (\<Sum>(x,K)\<in>{(x, \<G> K) |x K. (x,K) \<in> p \<and> \<G> K \<noteq> {}}. content K *\<^sub>R f x) =
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1217 | (\<Sum>(x,K)\<in>(\<lambda>(x,K). (x, \<G> K)) ` p. content K *\<^sub>R f x)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1218 | by (rule sum.mono_neutral_left) (auto simp: \<open>finite p\<close>) | 
| 53468 | 1219 | have "((\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) - i) + ((\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - j) = | 
| 1220 | (\<Sum>(x, k)\<in>?M1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>?M2. content k *\<^sub>R f x) - (i + j)" | |
| 1221 | by auto | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1222 |       moreover have "\<dots> = (\<Sum>(x,K) \<in> p. content (K \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f x) +
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1223 |         (\<Sum>(x,K) \<in> p. content (K \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f x) - (i + j)"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1224 | unfolding * | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1225 | apply (subst (1 2) sum.reindex_nontrivial) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1226 | apply (auto intro!: k p eq0 tagged_division_split_left_inj_content tagged_division_split_right_inj_content | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1227 | simp: cont_eq \<open>finite p\<close>) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1228 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1229 |       moreover have "\<And>x. x \<in> p \<Longrightarrow> (\<lambda>(a,B). content (B \<inter> {a. a \<bullet> k \<le> c}) *\<^sub>R f a) x +
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1230 |                                 (\<lambda>(a,B). content (B \<inter> {a. c \<le> a \<bullet> k}) *\<^sub>R f a) x =
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1231 | (\<lambda>(a,B). content B *\<^sub>R f a) x" | 
| 60435 | 1232 | proof clarify | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1233 | fix a B | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1234 | assume "(a, B) \<in> p" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1235 | with p obtain u v where uv: "B = cbox u v" by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1236 |         then show "content (B \<inter> {x. x \<bullet> k \<le> c}) *\<^sub>R f a + content (B \<inter> {x. c \<le> x \<bullet> k}) *\<^sub>R f a = content B *\<^sub>R f a"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1237 | by (auto simp: scaleR_left_distrib uv content_split[OF k,of u v c]) | 
| 53468 | 1238 | qed | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1239 | ultimately show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1240 | by (auto simp: sum.distrib[symmetric]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1241 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1242 | ultimately show "norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - (i + j)) < e" | 
| 53468 | 1243 | by auto | 
| 1244 | qed | |
| 1245 | qed | |
| 1246 | ||
| 35172 | 1247 | |
| 60420 | 1248 | subsection \<open>A sort of converse, integrability on subintervals.\<close> | 
| 35172 | 1249 | |
| 53494 | 1250 | lemma has_integral_separate_sides: | 
| 56188 | 1251 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 66359 | 1252 | assumes f: "(f has_integral i) (cbox a b)" | 
| 53494 | 1253 | and "e > 0" | 
| 1254 | and k: "k \<in> Basis" | |
| 1255 | obtains d where "gauge d" | |
| 56188 | 1256 |     "\<forall>p1 p2. p1 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<le> c}) \<and> d fine p1 \<and>
 | 
| 1257 |         p2 tagged_division_of (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) \<and> d fine p2 \<longrightarrow>
 | |
| 64267 | 1258 | norm ((sum (\<lambda>(x,k). content k *\<^sub>R f x) p1 + sum (\<lambda>(x,k). content k *\<^sub>R f x) p2) - i) < e" | 
| 53494 | 1259 | proof - | 
| 66359 | 1260 | obtain \<gamma> where d: "gauge \<gamma>" | 
| 1261 | "\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma> fine p\<rbrakk> | |
| 1262 | \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) < e" | |
| 1263 | using has_integralD[OF f \<open>e > 0\<close>] by metis | |
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1264 |   { fix p1 p2
 | 
| 66359 | 1265 |     assume tdiv1: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" and "\<gamma> fine p1"
 | 
| 1266 | note p1=tagged_division_ofD[OF this(1)] | |
| 1267 |     assume tdiv2: "p2 tagged_division_of (cbox a b) \<inter> {x. c \<le> x \<bullet> k}" and "\<gamma> fine p2"
 | |
| 1268 | note p2=tagged_division_ofD[OF this(1)] | |
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 1269 | note tagged_division_Un_interval[OF tdiv1 tdiv2] | 
| 66359 | 1270 | note p12 = tagged_division_ofD[OF this] this | 
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1271 |     { fix a b
 | 
| 53494 | 1272 | assume ab: "(a, b) \<in> p1 \<inter> p2" | 
| 1273 | have "(a, b) \<in> p1" | |
| 1274 | using ab by auto | |
| 66359 | 1275 | obtain u v where uv: "b = cbox u v" | 
| 1276 | using \<open>(a, b) \<in> p1\<close> p1(4) by moura | |
| 53494 | 1277 |       have "b \<subseteq> {x. x\<bullet>k = c}"
 | 
| 1278 | using ab p1(3)[of a b] p2(3)[of a b] by fastforce | |
| 1279 | moreover | |
| 1280 |       have "interior {x::'a. x \<bullet> k = c} = {}"
 | |
| 1281 | proof (rule ccontr) | |
| 1282 | assume "\<not> ?thesis" | |
| 1283 |         then obtain x where x: "x \<in> interior {x::'a. x\<bullet>k = c}"
 | |
| 1284 | by auto | |
| 66359 | 1285 |         then obtain \<epsilon> where "0 < \<epsilon>" and \<epsilon>: "ball x \<epsilon> \<subseteq> {x. x \<bullet> k = c}"
 | 
| 1286 | using mem_interior by metis | |
| 53494 | 1287 | have x: "x\<bullet>k = c" | 
| 1288 | using x interior_subset by fastforce | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 1289 | have *: "\<And>i. i \<in> Basis \<Longrightarrow> \<bar>(x - (x + (\<epsilon>/2) *\<^sub>R k)) \<bullet> i\<bar> = (if i = k then \<epsilon>/2 else 0)" | 
| 66359 | 1290 | using \<open>0 < \<epsilon>\<close> k by (auto simp: inner_simps inner_not_same_Basis) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 1291 | have "(\<Sum>i\<in>Basis. \<bar>(x - (x + (\<epsilon>/2 ) *\<^sub>R k)) \<bullet> i\<bar>) = | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 1292 | (\<Sum>i\<in>Basis. (if i = k then \<epsilon>/2 else 0))" | 
| 64267 | 1293 | using "*" by (blast intro: sum.cong) | 
| 66359 | 1294 | also have "\<dots> < \<epsilon>" | 
| 1295 | by (subst sum.delta) (use \<open>0 < \<epsilon>\<close> in auto) | |
| 1296 | finally have "x + (\<epsilon>/2) *\<^sub>R k \<in> ball x \<epsilon>" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1297 | unfolding mem_ball dist_norm by(rule le_less_trans[OF norm_le_l1]) | 
| 66359 | 1298 |         then have "x + (\<epsilon>/2) *\<^sub>R k \<in> {x. x\<bullet>k = c}"
 | 
| 1299 | using \<epsilon> by auto | |
| 53494 | 1300 | then show False | 
| 66359 | 1301 | using \<open>0 < \<epsilon>\<close> x k by (auto simp: inner_simps) | 
| 53494 | 1302 | qed | 
| 1303 | ultimately have "content b = 0" | |
| 1304 | unfolding uv content_eq_0_interior | |
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1305 | using interior_mono by blast | 
| 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1306 | then have "content b *\<^sub>R f a = 0" | 
| 53494 | 1307 | by auto | 
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1308 | } | 
| 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1309 | then have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) = | 
| 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1310 | norm ((\<Sum>(x, k)\<in>p1 \<union> p2. content k *\<^sub>R f x) - i)" | 
| 64267 | 1311 | by (subst sum.union_inter_neutral) (auto simp: p1 p2) | 
| 53494 | 1312 | also have "\<dots> < e" | 
| 66359 | 1313 | using d(2) p12 by (simp add: fine_Un k \<open>\<gamma> fine p1\<close> \<open>\<gamma> fine p2\<close>) | 
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1314 | finally have "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) + (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - i) < e" . | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 1315 | } | 
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1316 | then show ?thesis | 
| 66359 | 1317 | using d(1) that by auto | 
| 53494 | 1318 | qed | 
| 35172 | 1319 | |
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1320 | lemma integrable_split [intro]: | 
| 56188 | 1321 |   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
 | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1322 | assumes f: "f integrable_on cbox a b" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1323 | and k: "k \<in> Basis" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1324 |     shows "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<le> c})"   (is ?thesis1)
 | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1325 |     and   "f integrable_on (cbox a b \<inter> {x. x\<bullet>k \<ge> c})"   (is ?thesis2)
 | 
| 53494 | 1326 | proof - | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1327 | obtain y where y: "(f has_integral y) (cbox a b)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1328 | using f by blast | 
| 63040 | 1329 | define a' where "a' = (\<Sum>i\<in>Basis. (if i = k then max (a\<bullet>k) c else a\<bullet>i)*\<^sub>R i)" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1330 | define b' where "b' = (\<Sum>i\<in>Basis. (if i = k then min (b\<bullet>k) c else b\<bullet>i)*\<^sub>R i)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1331 | have "\<exists>d. gauge d \<and> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1332 |             (\<forall>p1 p2. p1 tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<le> c} \<and> d fine p1 \<and>
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1333 |                      p2 tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<le> c} \<and> d fine p2 \<longrightarrow>
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1334 | norm ((\<Sum>(x,K) \<in> p1. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> p2. content K *\<^sub>R f x)) < e)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1335 | if "e > 0" for e | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1336 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1337 | have "e/2 > 0" using that by auto | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1338 | with has_integral_separate_sides[OF y this k, of c] | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1339 | obtain d | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1340 | where "gauge d" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1341 |          and d: "\<And>p1 p2. \<lbrakk>p1 tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<le> c}; d fine p1;
 | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1342 |                           p2 tagged_division_of cbox a b \<inter> {x. c \<le> x \<bullet> k}; d fine p2\<rbrakk>
 | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1343 | \<Longrightarrow> norm ((\<Sum>(x,K)\<in>p1. content K *\<^sub>R f x) + (\<Sum>(x,K)\<in>p2. content K *\<^sub>R f x) - y) < e/2" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1344 | by metis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1345 | show ?thesis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1346 | proof (rule_tac x=d in exI, clarsimp simp add: \<open>gauge d\<close>) | 
| 53494 | 1347 | fix p1 p2 | 
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1348 |       assume as: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p1"
 | 
| 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1349 |                  "p2 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<le> c}" "d fine p2"
 | 
| 35172 | 1350 | show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1351 | proof (rule fine_division_exists[OF \<open>gauge d\<close>, of a' b]) | 
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1352 | fix p | 
| 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1353 | assume "p tagged_division_of cbox a' b" "d fine p" | 
| 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1354 | then show ?thesis | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1355 | using as norm_triangle_half_l[OF d[of p1 p] d[of p2 p]] | 
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1356 | unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] | 
| 53494 | 1357 | by (auto simp add: algebra_simps) | 
| 1358 | qed | |
| 1359 | qed | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1360 | qed | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1361 | with f show ?thesis1 | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1362 | by (simp add: interval_split[OF k] integrable_Cauchy) | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1363 | have "\<exists>d. gauge d \<and> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1364 |             (\<forall>p1 p2. p1 tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<ge> c} \<and> d fine p1 \<and>
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1365 |                      p2 tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<ge> c} \<and> d fine p2 \<longrightarrow>
 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1366 | norm ((\<Sum>(x,K) \<in> p1. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> p2. content K *\<^sub>R f x)) < e)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1367 | if "e > 0" for e | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1368 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1369 | have "e/2 > 0" using that by auto | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1370 | with has_integral_separate_sides[OF y this k, of c] | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1371 | obtain d | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1372 | where "gauge d" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1373 |          and d: "\<And>p1 p2. \<lbrakk>p1 tagged_division_of cbox a b \<inter> {x. x \<bullet> k \<le> c}; d fine p1;
 | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1374 |                           p2 tagged_division_of cbox a b \<inter> {x. c \<le> x \<bullet> k}; d fine p2\<rbrakk>
 | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1375 | \<Longrightarrow> norm ((\<Sum>(x,K)\<in>p1. content K *\<^sub>R f x) + (\<Sum>(x,K)\<in>p2. content K *\<^sub>R f x) - y) < e/2" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1376 | by metis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1377 | show ?thesis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1378 | proof (rule_tac x=d in exI, clarsimp simp add: \<open>gauge d\<close>) | 
| 53494 | 1379 | fix p1 p2 | 
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1380 |       assume as: "p1 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<ge> c}" "d fine p1"
 | 
| 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1381 |                  "p2 tagged_division_of (cbox a b) \<inter> {x. x \<bullet> k \<ge> c}" "d fine p2"
 | 
| 35172 | 1382 | show "norm ((\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)) < e" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1383 | proof (rule fine_division_exists[OF \<open>gauge d\<close>, of a b']) | 
| 60428 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1384 | fix p | 
| 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1385 | assume "p tagged_division_of cbox a b'" "d fine p" | 
| 
5e9de4faef98
fixed several "inside-out" proofs
 paulson <lp15@cam.ac.uk> parents: 
60425diff
changeset | 1386 | then show ?thesis | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1387 | using as norm_triangle_half_l[OF d[of p p1] d[of p p2]] | 
| 53494 | 1388 | unfolding interval_split[OF k] b'_def[symmetric] a'_def[symmetric] | 
| 53520 | 1389 | by (auto simp add: algebra_simps) | 
| 53494 | 1390 | qed | 
| 1391 | qed | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 1392 | qed | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 1393 | with f show ?thesis2 | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1394 | by (simp add: interval_split[OF k] integrable_Cauchy) | 
| 53494 | 1395 | qed | 
| 1396 | ||
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 1397 | lemma operative_integralI: | 
| 56188 | 1398 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | 
| 67399 | 1399 | shows "operative (lift_option (+)) (Some 0) | 
| 63659 | 1400 | (\<lambda>i. if f integrable_on i then Some (integral i f) else None)" | 
| 1401 | proof - | |
| 1402 | interpret comm_monoid "lift_option plus" "Some (0::'b)" | |
| 1403 | by (rule comm_monoid_lift_option) | |
| 1404 | (rule add.comm_monoid_axioms) | |
| 1405 | show ?thesis | |
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 1406 | proof | 
| 63659 | 1407 | fix a b c | 
| 1408 | fix k :: 'a | |
| 1409 | assume k: "k \<in> Basis" | |
| 1410 | show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = | |
| 67399 | 1411 |           lift_option (+) (if f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c} then Some (integral (cbox a b \<inter> {x. x \<bullet> k \<le> c}) f) else None)
 | 
| 63659 | 1412 |           (if f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k} then Some (integral (cbox a b \<inter> {x. c \<le> x \<bullet> k}) f) else None)"
 | 
| 1413 | proof (cases "f integrable_on cbox a b") | |
| 1414 | case True | |
| 1415 | with k show ?thesis | |
| 1416 | apply (simp add: integrable_split) | |
| 1417 | apply (rule integral_unique [OF has_integral_split[OF _ _ k]]) | |
| 60440 | 1418 | apply (auto intro: integrable_integral) | 
| 53494 | 1419 | done | 
| 63659 | 1420 | next | 
| 1421 | case False | |
| 1422 |       have "\<not> (f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}) \<or> \<not> ( f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k})"
 | |
| 1423 | proof (rule ccontr) | |
| 1424 | assume "\<not> ?thesis" | |
| 1425 | then have "f integrable_on cbox a b" | |
| 1426 | unfolding integrable_on_def | |
| 1427 |           apply (rule_tac x="integral (cbox a b \<inter> {x. x \<bullet> k \<le> c}) f + integral (cbox a b \<inter> {x. x \<bullet> k \<ge> c}) f" in exI)
 | |
| 1428 | apply (rule has_integral_split[OF _ _ k]) | |
| 1429 | apply (auto intro: integrable_integral) | |
| 1430 | done | |
| 1431 | then show False | |
| 1432 | using False by auto | |
| 1433 | qed | |
| 1434 | then show ?thesis | |
| 53494 | 1435 | using False by auto | 
| 1436 | qed | |
| 63659 | 1437 | next | 
| 1438 | fix a b :: 'a | |
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 1439 |     assume "box a b = {}"
 | 
| 63659 | 1440 | then show "(if f integrable_on cbox a b then Some (integral (cbox a b) f) else None) = Some 0" | 
| 1441 | using has_integral_null_eq | |
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 1442 | by (auto simp: integrable_on_null content_eq_0_interior) | 
| 63659 | 1443 | qed | 
| 53494 | 1444 | qed | 
| 1445 | ||
| 60420 | 1446 | subsection \<open>Bounds on the norm of Riemann sums and the integral itself.\<close> | 
| 35172 | 1447 | |
| 53494 | 1448 | lemma dsum_bound: | 
| 56188 | 1449 | assumes "p division_of (cbox a b)" | 
| 53494 | 1450 | and "norm c \<le> e" | 
| 64267 | 1451 | shows "norm (sum (\<lambda>l. content l *\<^sub>R c) p) \<le> e * content(cbox a b)" | 
| 60467 | 1452 | proof - | 
| 64267 | 1453 | have sumeq: "(\<Sum>i\<in>p. \<bar>content i\<bar>) = sum content p" | 
| 1454 | apply (rule sum.cong) | |
| 60467 | 1455 | using assms | 
| 1456 | apply simp | |
| 1457 | apply (metis abs_of_nonneg assms(1) content_pos_le division_ofD(4)) | |
| 1458 | done | |
| 1459 | have e: "0 \<le> e" | |
| 1460 | using assms(2) norm_ge_zero order_trans by blast | |
| 64267 | 1461 | have "norm (sum (\<lambda>l. content l *\<^sub>R c) p) \<le> (\<Sum>i\<in>p. norm (content i *\<^sub>R c))" | 
| 1462 | using norm_sum by blast | |
| 60467 | 1463 | also have "... \<le> e * (\<Sum>i\<in>p. \<bar>content i\<bar>)" | 
| 64267 | 1464 | by (simp add: sum_distrib_left[symmetric] mult.commute assms(2) mult_right_mono sum_nonneg) | 
| 60467 | 1465 | also have "... \<le> e * content (cbox a b)" | 
| 1466 | apply (rule mult_left_mono [OF _ e]) | |
| 1467 | apply (simp add: sumeq) | |
| 1468 | using additive_content_division assms(1) eq_iff apply blast | |
| 1469 | done | |
| 1470 | finally show ?thesis . | |
| 1471 | qed | |
| 53494 | 1472 | |
| 1473 | lemma rsum_bound: | |
| 60472 | 1474 | assumes p: "p tagged_division_of (cbox a b)" | 
| 1475 | and "\<forall>x\<in>cbox a b. norm (f x) \<le> e" | |
| 64267 | 1476 | shows "norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> e * content (cbox a b)" | 
| 56188 | 1477 | proof (cases "cbox a b = {}")
 | 
| 60472 | 1478 | case True show ?thesis | 
| 1479 | using p unfolding True tagged_division_of_trivial by auto | |
| 53494 | 1480 | next | 
| 1481 | case False | |
| 60472 | 1482 | then have e: "e \<ge> 0" | 
| 63018 
ae2ec7d86ad4
tidying some proofs; getting rid of "nonempty_witness"
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 1483 | by (meson ex_in_conv assms(2) norm_ge_zero order_trans) | 
| 64267 | 1484 | have sum_le: "sum (content \<circ> snd) p \<le> content (cbox a b)" | 
| 60472 | 1485 | unfolding additive_content_tagged_division[OF p, symmetric] split_def | 
| 1486 | by (auto intro: eq_refl) | |
| 1487 | have con: "\<And>xk. xk \<in> p \<Longrightarrow> 0 \<le> content (snd xk)" | |
| 1488 | using tagged_division_ofD(4) [OF p] content_pos_le | |
| 1489 | by force | |
| 1490 | have norm: "\<And>xk. xk \<in> p \<Longrightarrow> norm (f (fst xk)) \<le> e" | |
| 1491 | unfolding fst_conv using tagged_division_ofD(2,3)[OF p] assms | |
| 1492 | by (metis prod.collapse subset_eq) | |
| 64267 | 1493 | have "norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p) \<le> (\<Sum>i\<in>p. norm (case i of (x, k) \<Rightarrow> content k *\<^sub>R f x))" | 
| 1494 | by (rule norm_sum) | |
| 60472 | 1495 | also have "... \<le> e * content (cbox a b)" | 
| 53494 | 1496 | unfolding split_def norm_scaleR | 
| 64267 | 1497 | apply (rule order_trans[OF sum_mono]) | 
| 53494 | 1498 | apply (rule mult_left_mono[OF _ abs_ge_zero, of _ e]) | 
| 60472 | 1499 | apply (metis norm) | 
| 64267 | 1500 | unfolding sum_distrib_right[symmetric] | 
| 1501 | using con sum_le | |
| 60472 | 1502 | apply (auto simp: mult.commute intro: mult_left_mono [OF _ e]) | 
| 1503 | done | |
| 1504 | finally show ?thesis . | |
| 53494 | 1505 | qed | 
| 35172 | 1506 | |
| 1507 | lemma rsum_diff_bound: | |
| 56188 | 1508 | assumes "p tagged_division_of (cbox a b)" | 
| 1509 | and "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" | |
| 64267 | 1510 | shows "norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - sum (\<lambda>(x,k). content k *\<^sub>R g x) p) \<le> | 
| 60472 | 1511 | e * content (cbox a b)" | 
| 53494 | 1512 | apply (rule order_trans[OF _ rsum_bound[OF assms]]) | 
| 64267 | 1513 | apply (simp add: split_def scaleR_diff_right sum_subtractf eq_refl) | 
| 53494 | 1514 | done | 
| 1515 | ||
| 1516 | lemma has_integral_bound: | |
| 56188 | 1517 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 53494 | 1518 | assumes "0 \<le> B" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1519 | and f: "(f has_integral i) (cbox a b)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1520 | and "\<And>x. x\<in>cbox a b \<Longrightarrow> norm (f x) \<le> B" | 
| 60472 | 1521 | shows "norm i \<le> B * content (cbox a b)" | 
| 1522 | proof (rule ccontr) | |
| 53494 | 1523 | assume "\<not> ?thesis" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1524 | then have "norm i - B * content (cbox a b) > 0" | 
| 53494 | 1525 | by auto | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1526 | with f[unfolded has_integral] | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1527 | obtain \<gamma> where "gauge \<gamma>" and \<gamma>: | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1528 | "\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma> fine p\<rbrakk> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1529 | \<Longrightarrow> norm ((\<Sum>(x, K)\<in>p. content K *\<^sub>R f x) - i) < norm i - B * content (cbox a b)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1530 | by metis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1531 | then obtain p where p: "p tagged_division_of cbox a b" and "\<gamma> fine p" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1532 | using fine_division_exists by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1533 | have "\<And>s B. norm s \<le> B \<Longrightarrow> \<not> norm (s - i) < norm i - B" | 
| 60472 | 1534 | unfolding not_less | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1535 | by (metis diff_left_mono dist_commute dist_norm norm_triangle_ineq2 order_trans) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1536 | then show False | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1537 | using \<gamma> [OF p \<open>\<gamma> fine p\<close>] rsum_bound[OF p] assms by metis | 
| 53494 | 1538 | qed | 
| 1539 | ||
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 1540 | corollary integrable_bound: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 1541 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 1542 | assumes "0 \<le> B" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 1543 | and "f integrable_on (cbox a b)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 1544 | and "\<And>x. x\<in>cbox a b \<Longrightarrow> norm (f x) \<le> B" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 1545 | shows "norm (integral (cbox a b) f) \<le> B * content (cbox a b)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 1546 | by (metis integrable_integral has_integral_bound assms) | 
| 56188 | 1547 | |
| 35172 | 1548 | |
| 60420 | 1549 | subsection \<open>Similar theorems about relationship among components.\<close> | 
| 35172 | 1550 | |
| 53494 | 1551 | lemma rsum_component_le: | 
| 56188 | 1552 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1553 | assumes p: "p tagged_division_of (cbox a b)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1554 | and "\<And>x. x \<in> cbox a b \<Longrightarrow> (f x)\<bullet>i \<le> (g x)\<bullet>i" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1555 | shows "(\<Sum>(x, K)\<in>p. content K *\<^sub>R f x) \<bullet> i \<le> (\<Sum>(x, K)\<in>p. content K *\<^sub>R g x) \<bullet> i" | 
| 64267 | 1556 | unfolding inner_sum_left | 
| 1557 | proof (rule sum_mono, clarify) | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1558 | fix x K | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1559 | assume ab: "(x, K) \<in> p" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1560 | with p obtain u v where K: "K = cbox u v" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1561 | by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1562 | then show "(content K *\<^sub>R f x) \<bullet> i \<le> (content K *\<^sub>R g x) \<bullet> i" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 1563 | by (metis ab assms inner_scaleR_left measure_nonneg mult_left_mono tag_in_interval) | 
| 53494 | 1564 | qed | 
| 35172 | 1565 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1566 | lemma has_integral_component_le: | 
| 56188 | 1567 | fixes f g :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1568 | assumes k: "k \<in> Basis" | 
| 66199 | 1569 | assumes "(f has_integral i) S" "(g has_integral j) S" | 
| 1570 | and f_le_g: "\<And>x. x \<in> S \<Longrightarrow> (f x)\<bullet>k \<le> (g x)\<bullet>k" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1571 | shows "i\<bullet>k \<le> j\<bullet>k" | 
| 50348 | 1572 | proof - | 
| 66199 | 1573 | have ik_le_jk: "i\<bullet>k \<le> j\<bullet>k" | 
| 61165 | 1574 | if f_i: "(f has_integral i) (cbox a b)" | 
| 1575 | and g_j: "(g has_integral j) (cbox a b)" | |
| 1576 | and le: "\<forall>x\<in>cbox a b. (f x)\<bullet>k \<le> (g x)\<bullet>k" | |
| 1577 | for a b i and j :: 'b and f g :: "'a \<Rightarrow> 'b" | |
| 50348 | 1578 | proof (rule ccontr) | 
| 61165 | 1579 | assume "\<not> ?thesis" | 
| 53494 | 1580 | then have *: "0 < (i\<bullet>k - j\<bullet>k) / 3" | 
| 1581 | by auto | |
| 66199 | 1582 | obtain \<gamma>1 where "gauge \<gamma>1" | 
| 1583 | and \<gamma>1: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma>1 fine p\<rbrakk> | |
| 1584 | \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) < (i \<bullet> k - j \<bullet> k) / 3" | |
| 1585 | using f_i[unfolded has_integral,rule_format,OF *] by fastforce | |
| 1586 | obtain \<gamma>2 where "gauge \<gamma>2" | |
| 1587 | and \<gamma>2: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; \<gamma>2 fine p\<rbrakk> | |
| 1588 | \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - j) < (i \<bullet> k - j \<bullet> k) / 3" | |
| 1589 | using g_j[unfolded has_integral,rule_format,OF *] by fastforce | |
| 1590 | obtain p where p: "p tagged_division_of cbox a b" and "\<gamma>1 fine p" "\<gamma>2 fine p" | |
| 1591 | using fine_division_exists[OF gauge_Int[OF \<open>gauge \<gamma>1\<close> \<open>gauge \<gamma>2\<close>], of a b] unfolding fine_Int | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 1592 | by metis | 
| 60474 | 1593 | then have "\<bar>((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3" | 
| 66199 | 1594 | "\<bar>((\<Sum>(x, k)\<in>p. content k *\<^sub>R g x) - j) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3" | 
| 1595 | using le_less_trans[OF Basis_le_norm[OF k]] k \<gamma>1 \<gamma>2 by metis+ | |
| 53494 | 1596 | then show False | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1597 | unfolding inner_simps | 
| 66199 | 1598 | using rsum_component_le[OF p] le | 
| 1599 | by (fastforce simp add: abs_real_def split: if_split_asm) | |
| 53494 | 1600 | qed | 
| 60474 | 1601 | show ?thesis | 
| 66199 | 1602 | proof (cases "\<exists>a b. S = cbox a b") | 
| 60474 | 1603 | case True | 
| 66199 | 1604 | with ik_le_jk assms show ?thesis | 
| 60474 | 1605 | by auto | 
| 1606 | next | |
| 1607 | case False | |
| 1608 | show ?thesis | |
| 1609 | proof (rule ccontr) | |
| 1610 | assume "\<not> i\<bullet>k \<le> j\<bullet>k" | |
| 1611 | then have ij: "(i\<bullet>k - j\<bullet>k) / 3 > 0" | |
| 1612 | by auto | |
| 66199 | 1613 | obtain B1 where "0 < B1" | 
| 1614 | and B1: "\<And>a b. ball 0 B1 \<subseteq> cbox a b \<Longrightarrow> | |
| 1615 | \<exists>z. ((\<lambda>x. if x \<in> S then f x else 0) has_integral z) (cbox a b) \<and> | |
| 1616 | norm (z - i) < (i \<bullet> k - j \<bullet> k) / 3" | |
| 1617 | using has_integral_altD[OF _ False ij] assms by blast | |
| 1618 | obtain B2 where "0 < B2" | |
| 1619 | and B2: "\<And>a b. ball 0 B2 \<subseteq> cbox a b \<Longrightarrow> | |
| 1620 | \<exists>z. ((\<lambda>x. if x \<in> S then g x else 0) has_integral z) (cbox a b) \<and> | |
| 1621 | norm (z - j) < (i \<bullet> k - j \<bullet> k) / 3" | |
| 1622 | using has_integral_altD[OF _ False ij] assms by blast | |
| 60474 | 1623 | have "bounded (ball 0 B1 \<union> ball (0::'a) B2)" | 
| 1624 | unfolding bounded_Un by(rule conjI bounded_ball)+ | |
| 66199 | 1625 | from bounded_subset_cbox[OF this] | 
| 1626 | obtain a b::'a where ab: "ball 0 B1 \<subseteq> cbox a b" "ball 0 B2 \<subseteq> cbox a b" | |
| 66193 | 1627 | by blast+ | 
| 66199 | 1628 | then obtain w1 w2 where int_w1: "((\<lambda>x. if x \<in> S then f x else 0) has_integral w1) (cbox a b)" | 
| 1629 | and norm_w1: "norm (w1 - i) < (i \<bullet> k - j \<bullet> k) / 3" | |
| 1630 | and int_w2: "((\<lambda>x. if x \<in> S then g x else 0) has_integral w2) (cbox a b)" | |
| 1631 | and norm_w2: "norm (w2 - j) < (i \<bullet> k - j \<bullet> k) / 3" | |
| 1632 | using B1 B2 by blast | |
| 60474 | 1633 | have *: "\<And>w1 w2 j i::real .\<bar>w1 - i\<bar> < (i - j) / 3 \<Longrightarrow> \<bar>w2 - j\<bar> < (i - j) / 3 \<Longrightarrow> w1 \<le> w2 \<Longrightarrow> False" | 
| 62390 | 1634 | by (simp add: abs_real_def split: if_split_asm) | 
| 66199 | 1635 | have "\<bar>(w1 - i) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3" | 
| 1636 | "\<bar>(w2 - j) \<bullet> k\<bar> < (i \<bullet> k - j \<bullet> k) / 3" | |
| 1637 | using Basis_le_norm k le_less_trans norm_w1 norm_w2 by blast+ | |
| 60474 | 1638 | moreover | 
| 1639 | have "w1\<bullet>k \<le> w2\<bullet>k" | |
| 66199 | 1640 | using ik_le_jk int_w1 int_w2 f_le_g by auto | 
| 60474 | 1641 | ultimately show False | 
| 1642 | unfolding inner_simps by(rule *) | |
| 1643 | qed | |
| 1644 | qed | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1645 | qed | 
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36899diff
changeset | 1646 | |
| 53494 | 1647 | lemma integral_component_le: | 
| 56188 | 1648 | fixes g f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 53494 | 1649 | assumes "k \<in> Basis" | 
| 66199 | 1650 | and "f integrable_on S" "g integrable_on S" | 
| 1651 | and "\<And>x. x \<in> S \<Longrightarrow> (f x)\<bullet>k \<le> (g x)\<bullet>k" | |
| 1652 | shows "(integral S f)\<bullet>k \<le> (integral S g)\<bullet>k" | |
| 53494 | 1653 | apply (rule has_integral_component_le) | 
| 1654 | using integrable_integral assms | |
| 1655 | apply auto | |
| 1656 | done | |
| 1657 | ||
| 1658 | lemma has_integral_component_nonneg: | |
| 56188 | 1659 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 53494 | 1660 | assumes "k \<in> Basis" | 
| 66199 | 1661 | and "(f has_integral i) S" | 
| 1662 | and "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> (f x)\<bullet>k" | |
| 53494 | 1663 | shows "0 \<le> i\<bullet>k" | 
| 1664 | using has_integral_component_le[OF assms(1) has_integral_0 assms(2)] | |
| 1665 | using assms(3-) | |
| 1666 | by auto | |
| 1667 | ||
| 1668 | lemma integral_component_nonneg: | |
| 56188 | 1669 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 53494 | 1670 | assumes "k \<in> Basis" | 
| 66199 | 1671 | and "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> (f x)\<bullet>k" | 
| 1672 | shows "0 \<le> (integral S f)\<bullet>k" | |
| 1673 | proof (cases "f integrable_on S") | |
| 62463 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 1674 | case True show ?thesis | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 1675 | apply (rule has_integral_component_nonneg) | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 1676 | using assms True | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 1677 | apply auto | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 1678 | done | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 1679 | next | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 1680 | case False then show ?thesis by (simp add: not_integrable_integral) | 
| 
547c5c6e66d4
the integral is 0 when otherwise it would be undefined (also for contour integrals)
 paulson <lp15@cam.ac.uk> parents: 
62390diff
changeset | 1681 | qed | 
| 53494 | 1682 | |
| 1683 | lemma has_integral_component_neg: | |
| 56188 | 1684 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 53494 | 1685 | assumes "k \<in> Basis" | 
| 66199 | 1686 | and "(f has_integral i) S" | 
| 1687 | and "\<And>x. x \<in> S \<Longrightarrow> (f x)\<bullet>k \<le> 0" | |
| 53494 | 1688 | shows "i\<bullet>k \<le> 0" | 
| 1689 | using has_integral_component_le[OF assms(1,2) has_integral_0] assms(2-) | |
| 1690 | by auto | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1691 | |
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1692 | lemma has_integral_component_lbound: | 
| 56188 | 1693 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 1694 | assumes "(f has_integral i) (cbox a b)" | |
| 1695 | and "\<forall>x\<in>cbox a b. B \<le> f(x)\<bullet>k" | |
| 53494 | 1696 | and "k \<in> Basis" | 
| 56188 | 1697 | shows "B * content (cbox a b) \<le> i\<bullet>k" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1698 | using has_integral_component_le[OF assms(3) has_integral_const assms(1),of "(\<Sum>i\<in>Basis. B *\<^sub>R i)::'b"] assms(2-) | 
| 53494 | 1699 | by (auto simp add: field_simps) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1700 | |
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1701 | lemma has_integral_component_ubound: | 
| 56188 | 1702 | fixes f::"'a::euclidean_space => 'b::euclidean_space" | 
| 1703 | assumes "(f has_integral i) (cbox a b)" | |
| 1704 | and "\<forall>x\<in>cbox a b. f x\<bullet>k \<le> B" | |
| 53494 | 1705 | and "k \<in> Basis" | 
| 56188 | 1706 | shows "i\<bullet>k \<le> B * content (cbox a b)" | 
| 53494 | 1707 | using has_integral_component_le[OF assms(3,1) has_integral_const, of "\<Sum>i\<in>Basis. B *\<^sub>R i"] assms(2-) | 
| 1708 | by (auto simp add: field_simps) | |
| 1709 | ||
| 1710 | lemma integral_component_lbound: | |
| 56188 | 1711 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 1712 | assumes "f integrable_on cbox a b" | |
| 1713 | and "\<forall>x\<in>cbox a b. B \<le> f(x)\<bullet>k" | |
| 53494 | 1714 | and "k \<in> Basis" | 
| 56188 | 1715 | shows "B * content (cbox a b) \<le> (integral(cbox a b) f)\<bullet>k" | 
| 53494 | 1716 | apply (rule has_integral_component_lbound) | 
| 1717 | using assms | |
| 1718 | unfolding has_integral_integral | |
| 1719 | apply auto | |
| 1720 | done | |
| 1721 | ||
| 56190 | 1722 | lemma integral_component_lbound_real: | 
| 66402 | 1723 |   assumes "f integrable_on {a ::real..b}"
 | 
| 1724 |     and "\<forall>x\<in>{a..b}. B \<le> f(x)\<bullet>k"
 | |
| 56190 | 1725 | and "k \<in> Basis" | 
| 66402 | 1726 |   shows "B * content {a..b} \<le> (integral {a..b} f)\<bullet>k"
 | 
| 56190 | 1727 | using assms | 
| 1728 | by (metis box_real(2) integral_component_lbound) | |
| 1729 | ||
| 53494 | 1730 | lemma integral_component_ubound: | 
| 56188 | 1731 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::euclidean_space" | 
| 1732 | assumes "f integrable_on cbox a b" | |
| 1733 | and "\<forall>x\<in>cbox a b. f x\<bullet>k \<le> B" | |
| 53494 | 1734 | and "k \<in> Basis" | 
| 56188 | 1735 | shows "(integral (cbox a b) f)\<bullet>k \<le> B * content (cbox a b)" | 
| 53494 | 1736 | apply (rule has_integral_component_ubound) | 
| 1737 | using assms | |
| 1738 | unfolding has_integral_integral | |
| 1739 | apply auto | |
| 1740 | done | |
| 1741 | ||
| 56190 | 1742 | lemma integral_component_ubound_real: | 
| 1743 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" | |
| 66402 | 1744 |   assumes "f integrable_on {a..b}"
 | 
| 1745 |     and "\<forall>x\<in>{a..b}. f x\<bullet>k \<le> B"
 | |
| 56190 | 1746 | and "k \<in> Basis" | 
| 66402 | 1747 |   shows "(integral {a..b} f)\<bullet>k \<le> B * content {a..b}"
 | 
| 56190 | 1748 | using assms | 
| 1749 | by (metis box_real(2) integral_component_ubound) | |
| 35172 | 1750 | |
| 60420 | 1751 | subsection \<open>Uniform limit of integrable functions is integrable.\<close> | 
| 35172 | 1752 | |
| 62626 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1753 | lemma real_arch_invD: | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1754 | "0 < (e::real) \<Longrightarrow> (\<exists>n::nat. n \<noteq> 0 \<and> 0 < inverse (real n) \<and> inverse (real n) < e)" | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1755 | by (subst(asm) real_arch_inverse) | 
| 
de25474ce728
Contractible sets. Also removal of obsolete theorems and refactoring
 paulson <lp15@cam.ac.uk> parents: 
62623diff
changeset | 1756 | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1757 | |
| 53494 | 1758 | lemma integrable_uniform_limit: | 
| 56188 | 1759 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1760 | assumes "\<And>e. e > 0 \<Longrightarrow> \<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b" | 
| 56188 | 1761 | shows "f integrable_on cbox a b" | 
| 60487 | 1762 | proof (cases "content (cbox a b) > 0") | 
| 1763 | case False then show ?thesis | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1764 | using has_integral_null by (simp add: content_lt_nz integrable_on_def) | 
| 60487 | 1765 | next | 
| 1766 | case True | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1767 | have "1 / (real n + 1) > 0" for n | 
| 53494 | 1768 | by auto | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1769 | then have "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> 1 / (real n + 1)) \<and> g integrable_on cbox a b" for n | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1770 | using assms by blast | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1771 | then obtain g where g_near_f: "\<And>n x. x \<in> cbox a b \<Longrightarrow> norm (f x - g n x) \<le> 1 / (real n + 1)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1772 | and int_g: "\<And>n. g n integrable_on cbox a b" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1773 | by metis | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1774 | then obtain h where h: "\<And>n. (g n has_integral h n) (cbox a b)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1775 | unfolding integrable_on_def by metis | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1776 | have "Cauchy h" | 
| 53494 | 1777 | unfolding Cauchy_def | 
| 60487 | 1778 | proof clarify | 
| 53494 | 1779 | fix e :: real | 
| 1780 | assume "e>0" | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1781 | then have "e/4 / content (cbox a b) > 0" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1782 | using True by (auto simp: field_simps) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1783 | then obtain M where "M \<noteq> 0" and M: "1 / (real M) < e/4 / content (cbox a b)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1784 | by (metis inverse_eq_divide real_arch_inverse) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1785 | show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (h m) (h n) < e" | 
| 60487 | 1786 | proof (rule exI [where x=M], clarify) | 
| 1787 | fix m n | |
| 1788 | assume m: "M \<le> m" and n: "M \<le> n" | |
| 60420 | 1789 | have "e/4>0" using \<open>e>0\<close> by auto | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1790 | then obtain gm gn where "gauge gm" "gauge gn" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1791 | and gm: "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> gm fine \<D> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1792 | \<Longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g m x) - h m) < e/4" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1793 | and gn: "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> gn fine \<D> \<Longrightarrow> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1794 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g n x) - h n) < e/4" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1795 | using h[unfolded has_integral] by meson | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1796 | then obtain \<D> where \<D>: "\<D> tagged_division_of cbox a b" "(\<lambda>x. gm x \<inter> gn x) fine \<D>" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1797 | by (metis (full_types) fine_division_exists gauge_Int) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1798 | have triangle3: "norm (i1 - i2) < e" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1799 | if no: "norm(s2 - s1) \<le> e/2" "norm (s1 - i1) < e/4" "norm (s2 - i2) < e/4" for s1 s2 i1 and i2::'b | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1800 | proof - | 
| 53494 | 1801 | have "norm (i1 - i2) \<le> norm (i1 - s1) + norm (s1 - s2) + norm (s2 - i2)" | 
| 35172 | 1802 | using norm_triangle_ineq[of "i1 - s1" "s1 - i2"] | 
| 53494 | 1803 | using norm_triangle_ineq[of "s1 - s2" "s2 - i2"] | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1804 | by (auto simp: algebra_simps) | 
| 53494 | 1805 | also have "\<dots> < e" | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1806 | using no by (auto simp: algebra_simps norm_minus_commute) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1807 | finally show ?thesis . | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1808 | qed | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1809 | have finep: "gm fine \<D>" "gn fine \<D>" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1810 | using fine_Int \<D> by auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1811 | have norm_le: "norm (g n x - g m x) \<le> 2 / real M" if x: "x \<in> cbox a b" for x | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1812 | proof - | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1813 | have "norm (f x - g n x) + norm (f x - g m x) \<le> 1 / (real n + 1) + 1 / (real m + 1)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1814 | using g_near_f[OF x, of n] g_near_f[OF x, of m] by simp | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1815 | also have "\<dots> \<le> 1 / (real M) + 1 / (real M)" | 
| 53494 | 1816 | apply (rule add_mono) | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1817 | using \<open>M \<noteq> 0\<close> m n by (auto simp: divide_simps) | 
| 53494 | 1818 | also have "\<dots> = 2 / real M" | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1819 | by auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1820 | finally show "norm (g n x - g m x) \<le> 2 / real M" | 
| 35172 | 1821 | using norm_triangle_le[of "g n x - f x" "f x - g m x" "2 / real M"] | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1822 | by (auto simp: algebra_simps simp add: norm_minus_commute) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1823 | qed | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1824 | have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g n x) - (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g m x)) \<le> 2 / real M * content (cbox a b)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1825 | by (blast intro: norm_le rsum_diff_bound[OF \<D>(1), where e="2 / real M"]) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1826 | also have "... \<le> e/2" | 
| 60487 | 1827 | using M True | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1828 | by (auto simp: field_simps) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1829 | finally have le_e2: "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g n x) - (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g m x)) \<le> e/2" . | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1830 | then show "dist (h m) (h n) < e" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1831 | unfolding dist_norm using gm gn \<D> finep by (auto intro!: triangle3) | 
| 60487 | 1832 | qed | 
| 1833 | qed | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1834 | then obtain s where s: "h \<longlonglongrightarrow> s" | 
| 64287 | 1835 | using convergent_eq_Cauchy[symmetric] by blast | 
| 53494 | 1836 | show ?thesis | 
| 60487 | 1837 | unfolding integrable_on_def has_integral | 
| 1838 | proof (rule_tac x=s in exI, clarify) | |
| 1839 | fix e::real | |
| 1840 | assume e: "0 < e" | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1841 | then have "e/3 > 0" by auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1842 | then obtain N1 where N1: "\<forall>n\<ge>N1. norm (h n - s) < e/3" | 
| 60487 | 1843 | using LIMSEQ_D [OF s] by metis | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1844 | from e True have "e/3 / content (cbox a b) > 0" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1845 | by (auto simp: field_simps) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1846 | then obtain N2 :: nat | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1847 | where "N2 \<noteq> 0" and N2: "1 / (real N2) < e/3 / content (cbox a b)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1848 | by (metis inverse_eq_divide real_arch_inverse) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1849 | obtain g' where "gauge g'" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1850 | and g': "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> g' fine \<D> \<Longrightarrow> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1851 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1852 | by (metis h has_integral \<open>e/3 > 0\<close>) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1853 | have *: "norm (sf - s) < e" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1854 | if no: "norm (sf - sg) \<le> e/3" "norm(h - s) < e/3" "norm (sg - h) < e/3" for sf sg h | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1855 | proof - | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1856 | have "norm (sf - s) \<le> norm (sf - sg) + norm (sg - h) + norm (h - s)" | 
| 35172 | 1857 | using norm_triangle_ineq[of "sf - sg" "sg - s"] | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1858 | using norm_triangle_ineq[of "sg - h" " h - s"] | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1859 | by (auto simp: algebra_simps) | 
| 53494 | 1860 | also have "\<dots> < e" | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1861 | using no by (auto simp: algebra_simps norm_minus_commute) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1862 | finally show ?thesis . | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1863 | qed | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1864 |     { fix \<D>
 | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1865 | assume ptag: "\<D> tagged_division_of (cbox a b)" and "g' fine \<D>" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1866 | then have norm_less: "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g (N1 + N2) x) - h (N1 + N2)) < e/3" | 
| 60487 | 1867 | using g' by blast | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1868 | have "content (cbox a b) < e/3 * (of_nat N2)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1869 | using \<open>N2 \<noteq> 0\<close> N2 using True by (auto simp: divide_simps) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1870 | moreover have "e/3 * of_nat N2 \<le> e/3 * (of_nat (N1 + N2) + 1)" | 
| 60487 | 1871 | using \<open>e>0\<close> by auto | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1872 | ultimately have "content (cbox a b) < e/3 * (of_nat (N1 + N2) + 1)" | 
| 60487 | 1873 | by linarith | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1874 | then have le_e3: "1 / (real (N1 + N2) + 1) * content (cbox a b) \<le> e/3" | 
| 60487 | 1875 | unfolding inverse_eq_divide | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1876 | by (auto simp: field_simps) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1877 | have ne3: "norm (h (N1 + N2) - s) < e/3" | 
| 60487 | 1878 | using N1 by auto | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1879 | have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R g (N1 + N2) x)) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1880 | \<le> 1 / (real (N1 + N2) + 1) * content (cbox a b)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1881 | by (blast intro: g_near_f rsum_diff_bound[OF ptag]) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1882 | then have "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - s) < e" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1883 | by (rule *[OF order_trans [OF _ le_e3] ne3 norm_less]) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1884 | } | 
| 60487 | 1885 | then show "\<exists>d. gauge d \<and> | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1886 | (\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> d fine \<D> \<longrightarrow> norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - s) < e)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1887 | by (blast intro: g' \<open>gauge g'\<close>) | 
| 53494 | 1888 | qed | 
| 1889 | qed | |
| 1890 | ||
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 1891 | lemmas integrable_uniform_limit_real = integrable_uniform_limit [where 'a=real, simplified] | 
| 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 1892 | |
| 35172 | 1893 | |
| 60420 | 1894 | subsection \<open>Negligible sets.\<close> | 
| 35172 | 1895 | |
| 56188 | 1896 | definition "negligible (s:: 'a::euclidean_space set) \<longleftrightarrow> | 
| 1897 | (\<forall>a b. ((indicator s :: 'a\<Rightarrow>real) has_integral 0) (cbox a b))" | |
| 53494 | 1898 | |
| 35172 | 1899 | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 1900 | subsubsection \<open>Negligibility of hyperplane.\<close> | 
| 35172 | 1901 | |
| 53495 | 1902 | lemma content_doublesplit: | 
| 56188 | 1903 | fixes a :: "'a::euclidean_space" | 
| 53495 | 1904 | assumes "0 < e" | 
| 1905 | and k: "k \<in> Basis" | |
| 61945 | 1906 |   obtains d where "0 < d" and "content (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) < e"
 | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1907 | proof cases | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1908 | assume *: "a \<bullet> k \<le> c \<and> c \<le> b \<bullet> k \<and> (\<forall>j\<in>Basis. a \<bullet> j \<le> b \<bullet> j)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1909 | define a' where "a' d = (\<Sum>j\<in>Basis. (if j = k then max (a\<bullet>j) (c - d) else a\<bullet>j) *\<^sub>R j)" for d | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1910 | define b' where "b' d = (\<Sum>j\<in>Basis. (if j = k then min (b\<bullet>j) (c + d) else b\<bullet>j) *\<^sub>R j)" for d | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1911 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1912 | have "((\<lambda>d. \<Prod>j\<in>Basis. (b' d - a' d) \<bullet> j) \<longlongrightarrow> (\<Prod>j\<in>Basis. (b' 0 - a' 0) \<bullet> j)) (at_right 0)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1913 | by (auto simp: b'_def a'_def intro!: tendsto_min tendsto_max tendsto_eq_intros) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1914 | also have "(\<Prod>j\<in>Basis. (b' 0 - a' 0) \<bullet> j) = 0" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1915 | using k * | 
| 64272 | 1916 | by (intro prod_zero bexI[OF _ k]) | 
| 64267 | 1917 | (auto simp: b'_def a'_def inner_diff inner_sum_left inner_not_same_Basis intro!: sum.cong) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1918 | also have "((\<lambda>d. \<Prod>j\<in>Basis. (b' d - a' d) \<bullet> j) \<longlongrightarrow> 0) (at_right 0) = | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1919 |     ((\<lambda>d. content (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d})) \<longlongrightarrow> 0) (at_right 0)"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1920 | proof (intro tendsto_cong eventually_at_rightI) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1921 |     fix d :: real assume d: "d \<in> {0<..<1}"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1922 |     have "cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d} = cbox (a' d) (b' d)" for d
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1923 | using * d k by (auto simp add: cbox_def set_eq_iff Int_def ball_conj_distrib abs_diff_le_iff a'_def b'_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1924 | moreover have "j \<in> Basis \<Longrightarrow> a' d \<bullet> j \<le> b' d \<bullet> j" for j | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1925 | using * d k by (auto simp: a'_def b'_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1926 |     ultimately show "(\<Prod>j\<in>Basis. (b' d - a' d) \<bullet> j) = content (cbox a b \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d})"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1927 | by simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1928 | qed simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1929 |   finally have "((\<lambda>d. content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) \<longlongrightarrow> 0) (at_right 0)" .
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1930 | from order_tendstoD(2)[OF this \<open>0<e\<close>] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1931 |   obtain d' where "0 < d'" and d': "\<And>y. y > 0 \<Longrightarrow> y < d' \<Longrightarrow> content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> y}) < e"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1932 | by (subst (asm) eventually_at_right[of _ 1]) auto | 
| 53495 | 1933 | show ?thesis | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1934 | by (rule that[of "d'/2"], insert \<open>0<d'\<close> d'[of "d'/2"], auto) | 
| 53495 | 1935 | next | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1936 | assume *: "\<not> (a \<bullet> k \<le> c \<and> c \<le> b \<bullet> k \<and> (\<forall>j\<in>Basis. a \<bullet> j \<le> b \<bullet> j))" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1937 | then have "(\<exists>j\<in>Basis. b \<bullet> j < a \<bullet> j) \<or> (c < a \<bullet> k \<or> b \<bullet> k < c)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1938 | by (auto simp: not_le) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1939 | show thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1940 | proof cases | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1941 | assume "\<exists>j\<in>Basis. b \<bullet> j < a \<bullet> j" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1942 |     then have [simp]: "cbox a b = {}"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1943 | using box_ne_empty(1)[of a b] by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1944 | show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1945 | by (rule that[of 1]) (simp_all add: \<open>0<e\<close>) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1946 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1947 | assume "\<not> (\<exists>j\<in>Basis. b \<bullet> j < a \<bullet> j)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1948 | with * have "c < a \<bullet> k \<or> b \<bullet> k < c" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1949 | by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1950 | then show thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1951 | proof | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1952 | assume c: "c < a \<bullet> k" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1953 | moreover have "x \<in> cbox a b \<Longrightarrow> c \<le> x \<bullet> k" for x | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1954 | using k c by (auto simp: cbox_def) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 1955 |       ultimately have "cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> (a \<bullet> k - c)/2} = {}"
 | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1956 | using k by (auto simp: cbox_def) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 1957 | with \<open>0<e\<close> c that[of "(a \<bullet> k - c)/2"] show ?thesis | 
| 53495 | 1958 | by auto | 
| 60492 | 1959 | next | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1960 | assume c: "b \<bullet> k < c" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1961 | moreover have "x \<in> cbox a b \<Longrightarrow> x \<bullet> k \<le> c" for x | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1962 | using k c by (auto simp: cbox_def) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 1963 |       ultimately have "cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> (c - b \<bullet> k)/2} = {}"
 | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1964 | using k by (auto simp: cbox_def) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 1965 | with \<open>0<e\<close> c that[of "(c - b \<bullet> k)/2"] show ?thesis | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1966 | by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1967 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1968 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1969 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 1970 | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1971 | |
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1972 | proposition negligible_standard_hyperplane[intro]: | 
| 56188 | 1973 | fixes k :: "'a::euclidean_space" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1974 | assumes k: "k \<in> Basis" | 
| 53399 | 1975 |   shows "negligible {x. x\<bullet>k = c}"
 | 
| 53495 | 1976 | unfolding negligible_def has_integral | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1977 | proof clarsimp | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1978 | fix a b and e::real assume "e > 0" | 
| 66537 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 1979 |   with k obtain d where "0 < d" and d: "content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) < e"
 | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1980 | by (metis content_doublesplit) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 1981 |   let ?i = "indicator {x::'a. x\<bullet>k = c} :: 'a\<Rightarrow>real"
 | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1982 | show "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1983 | (\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1984 | \<bar>\<Sum>(x,K) \<in> \<D>. content K * ?i x\<bar> < e)" | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1985 | proof (intro exI, safe) | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1986 | show "gauge (\<lambda>x. ball x d)" | 
| 66537 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 1987 | using \<open>0 < d\<close> by blast | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1988 | next | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1989 | fix \<D> | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 1990 | assume p: "\<D> tagged_division_of (cbox a b)" "(\<lambda>x. ball x d) fine \<D>" | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 1991 |     have "content L = content (L \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})" 
 | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 1992 | if "(x, L) \<in> \<D>" "?i x \<noteq> 0" for x L | 
| 53495 | 1993 | proof - | 
| 66537 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 1994 | have xk: "x\<bullet>k = c" | 
| 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 1995 | using that by (simp add: indicator_def split: if_split_asm) | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 1996 |       have "L \<subseteq> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}"
 | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 1997 | proof | 
| 53495 | 1998 | fix y | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 1999 | assume y: "y \<in> L" | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2000 | have "L \<subseteq> ball x d" | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2001 | using p(2) that(1) by auto | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2002 | then have "norm (x - y) < d" | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2003 | by (simp add: dist_norm subset_iff y) | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2004 | then have "\<bar>(x - y) \<bullet> k\<bar> < d" | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2005 | using k norm_bound_Basis_lt by blast | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2006 |         then show "y \<in> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}"
 | 
| 53495 | 2007 | unfolding inner_simps xk by auto | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2008 | qed | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2009 |       then show "content L = content (L \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})"
 | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2010 | by (metis inf.orderE) | 
| 53495 | 2011 | qed | 
| 66537 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 2012 |     then have *: "(\<Sum>(x,K)\<in>\<D>. content K * ?i x) = (\<Sum>(x,K)\<in>\<D>. content (K \<inter> {x. \<bar>x\<bullet>k - c\<bar> \<le> d}) *\<^sub>R ?i x)"
 | 
| 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 2013 | by (force simp add: split_paired_all intro!: sum.cong [OF refl]) | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2014 | note p'= tagged_division_ofD[OF p(1)] and p''=division_of_tagged_division[OF p(1)] | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2015 |     have "(\<Sum>(x,K)\<in>\<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * indicator {x. x \<bullet> k = c} x) < e"
 | 
| 53495 | 2016 | proof - | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2017 |       have "(\<Sum>(x,K)\<in>\<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) \<le> (\<Sum>(x,K)\<in>\<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}))"
 | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2018 | by (force simp add: indicator_def intro!: sum_mono) | 
| 53495 | 2019 | also have "\<dots> < e" | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2020 | proof (subst sum.over_tagged_division_lemma[OF p(1)]) | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2021 | fix u v::'a | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2022 |         assume "box u v = {}"
 | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 2023 | then have *: "content (cbox u v) = 0" | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 2024 | unfolding content_eq_0_interior by simp | 
| 66537 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 2025 |         have "cbox u v \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<subseteq> cbox u v"
 | 
| 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 2026 | by auto | 
| 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 2027 |         then have "content (cbox u v \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<le> content (cbox u v)"
 | 
| 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 2028 | unfolding interval_doublesplit[OF k] by (rule content_subset) | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2029 |         then show "content (cbox u v \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = 0"
 | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 2030 | unfolding * interval_doublesplit[OF k] | 
| 50348 | 2031 | by (blast intro: antisym) | 
| 53495 | 2032 | next | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2033 |         have "(\<Sum>l\<in>snd ` \<D>. content (l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) =
 | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2034 |           sum content ((\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` \<D>. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}})"
 | 
| 64267 | 2035 | proof (subst (2) sum.reindex_nontrivial) | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2036 |           fix x y assume "x \<in> {l \<in> snd ` \<D>. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}" "y \<in> {l \<in> snd ` \<D>. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}"
 | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2037 |             "x \<noteq> y" and eq: "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} = y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}"
 | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2038 |           then obtain x' y' where "(x', x) \<in> \<D>" "x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}" "(y', y) \<in> \<D>" "y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}"
 | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2039 | by (auto) | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2040 |           from p'(5)[OF \<open>(x', x) \<in> \<D>\<close> \<open>(y', y) \<in> \<D>\<close>] \<open>x \<noteq> y\<close> have "interior (x \<inter> y) = {}"
 | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2041 | by auto | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2042 |           moreover have "interior ((x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) \<inter> (y \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) \<subseteq> interior (x \<inter> y)"
 | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2043 | by (auto intro: interior_mono) | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2044 |           ultimately have "interior (x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = {}"
 | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2045 | by (auto simp: eq) | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2046 |           then show "content (x \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) = 0"
 | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2047 | using p'(4)[OF \<open>(x', x) \<in> \<D>\<close>] by (auto simp: interval_doublesplit[OF k] content_eq_0_interior simp del: interior_Int) | 
| 64267 | 2048 | qed (insert p'(1), auto intro!: sum.mono_neutral_right) | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2049 |         also have "\<dots> \<le> norm (\<Sum>l\<in>(\<lambda>l. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})`{l\<in>snd ` \<D>. l \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d} \<noteq> {}}. content l *\<^sub>R 1::real)"
 | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2050 | by simp | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2051 |         also have "\<dots> \<le> 1 * content (cbox a b \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})"
 | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2052 | using division_doublesplit[OF p'' k, unfolded interval_doublesplit[OF k]] | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2053 | unfolding interval_doublesplit[OF k] by (intro dsum_bound) auto | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2054 | also have "\<dots> < e" | 
| 66537 
e2249cd6df67
sorted out cases in negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66536diff
changeset | 2055 | using d by simp | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2056 |         finally show "(\<Sum>K\<in>snd ` \<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d})) < e" .
 | 
| 53495 | 2057 | qed | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2058 |       finally show "(\<Sum>(x, K)\<in>\<D>. content (K \<inter> {x. \<bar>x \<bullet> k - c\<bar> \<le> d}) * ?i x) < e" .
 | 
| 53495 | 2059 | qed | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2060 | then show "\<bar>\<Sum>(x, K)\<in>\<D>. content K * ?i x\<bar> < e" | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2061 | unfolding * | 
| 65680 
378a2f11bec9
Simplification of some proofs. Also key lemmas using !! rather than ! in premises
 paulson <lp15@cam.ac.uk> parents: 
65587diff
changeset | 2062 | apply (subst abs_of_nonneg) | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2063 | using measure_nonneg | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2064 | by (force simp add: indicator_def intro: sum_nonneg)+ | 
| 53495 | 2065 | qed | 
| 2066 | qed | |
| 2067 | ||
| 35172 | 2068 | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2069 | subsubsection \<open>Hence the main theorem about negligible sets.\<close> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2070 | |
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2071 | |
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2072 | lemma has_integral_negligible_cbox: | 
| 56188 | 2073 | fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector" | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2074 | assumes negs: "negligible S" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2075 | and 0: "\<And>x. x \<notin> S \<Longrightarrow> f x = 0" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2076 | shows "(f has_integral 0) (cbox a b)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2077 | unfolding has_integral | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2078 | proof clarify | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2079 | fix e::real | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2080 | assume "e > 0" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2081 | then have nn_gt0: "e/2 / ((real n+1) * (2 ^ n)) > 0" for n | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2082 | by simp | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2083 | then have "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2084 | (\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2085 | \<bar>\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R indicator S x\<bar> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2086 | < e/2 / ((real n + 1) * 2 ^ n))" for n | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2087 | using negs [unfolded negligible_def has_integral] by auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2088 | then obtain \<gamma> where | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2089 | gd: "\<And>n. gauge (\<gamma> n)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2090 | and \<gamma>: "\<And>n \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> n fine \<D>\<rbrakk> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2091 | \<Longrightarrow> \<bar>\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R indicator S x\<bar> < e/2 / ((real n + 1) * 2 ^ n)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2092 | by metis | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2093 | show "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2094 | (\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2095 | norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - 0) < e)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2096 | proof (intro exI, safe) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2097 | show "gauge (\<lambda>x. \<gamma> (nat \<lfloor>norm (f x)\<rfloor>) x)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2098 | using gd by (auto simp: gauge_def) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2099 | |
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2100 | show "norm ((\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) - 0) < e" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2101 | if "\<D> tagged_division_of (cbox a b)" "(\<lambda>x. \<gamma> (nat \<lfloor>norm (f x)\<rfloor>) x) fine \<D>" for \<D> | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2102 |     proof (cases "\<D> = {}")
 | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2103 | case True with \<open>0 < e\<close> show ?thesis by simp | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2104 | next | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2105 | case False | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2106 | obtain N where "Max ((\<lambda>(x, K). norm (f x)) ` \<D>) \<le> real N" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2107 | using real_arch_simple by blast | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2108 | then have N: "\<And>x. x \<in> (\<lambda>(x, K). norm (f x)) ` \<D> \<Longrightarrow> x \<le> real N" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2109 | by (meson Max_ge that(1) dual_order.trans finite_imageI tagged_division_of_finite) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2110 | have "\<forall>i. \<exists>q. q tagged_division_of (cbox a b) \<and> (\<gamma> i) fine q \<and> (\<forall>(x,K) \<in> \<D>. K \<subseteq> (\<gamma> i) x \<longrightarrow> (x, K) \<in> q)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2111 | by (auto intro: tagged_division_finer[OF that(1) gd]) | 
| 66199 | 2112 | from choice[OF this] | 
| 2113 | obtain q where q: "\<And>n. q n tagged_division_of cbox a b" | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2114 | "\<And>n. \<gamma> n fine q n" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2115 | "\<And>n x K. \<lbrakk>(x, K) \<in> \<D>; K \<subseteq> \<gamma> n x\<rbrakk> \<Longrightarrow> (x, K) \<in> q n" | 
| 66199 | 2116 | by fastforce | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2117 | have "finite \<D>" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2118 | using that(1) by blast | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2119 | then have sum_le_inc: "\<lbrakk>finite T; \<And>x y. (x,y) \<in> T \<Longrightarrow> (0::real) \<le> g(x,y); | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2120 | \<And>y. y\<in>\<D> \<Longrightarrow> \<exists>x. (x,y) \<in> T \<and> f(y) \<le> g(x,y)\<rbrakk> \<Longrightarrow> sum f \<D> \<le> sum g T" for f g T | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2121 | by (rule sum_le_included[of \<D> T g snd f]; force) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2122 | have "norm (\<Sum>(x,K) \<in> \<D>. content K *\<^sub>R f x) \<le> (\<Sum>(x,K) \<in> \<D>. norm (content K *\<^sub>R f x))" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2123 | unfolding split_def by (rule norm_sum) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2124 |       also have "... \<le> (\<Sum>(i, j) \<in> Sigma {..N + 1} q.
 | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2125 | (real i + 1) * (case j of (x, K) \<Rightarrow> content K *\<^sub>R indicator S x))" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2126 | proof (rule sum_le_inc, safe) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2127 |         show "finite (Sigma {..N+1} q)"
 | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2128 | by (meson finite_SigmaI finite_atMost tagged_division_of_finite q(1)) | 
| 53495 | 2129 | next | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2130 | fix x K | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2131 | assume xk: "(x, K) \<in> \<D>" | 
| 63040 | 2132 | define n where "n = nat \<lfloor>norm (f x)\<rfloor>" | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2133 | have *: "norm (f x) \<in> (\<lambda>(x, K). norm (f x)) ` \<D>" | 
| 53495 | 2134 | using xk by auto | 
| 2135 | have nfx: "real n \<le> norm (f x)" "norm (f x) \<le> real n + 1" | |
| 2136 | unfolding n_def by auto | |
| 2137 |         then have "n \<in> {0..N + 1}"
 | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2138 | using N[OF *] by auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2139 | moreover have "K \<subseteq> \<gamma> (nat \<lfloor>norm (f x)\<rfloor>) x" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2140 | using that(2) xk by auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2141 | moreover then have "(x, K) \<in> q (nat \<lfloor>norm (f x)\<rfloor>)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2142 | by (simp add: q(3) xk) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2143 | moreover then have "(x, K) \<in> q n" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2144 | using n_def by blast | 
| 53495 | 2145 | moreover | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2146 | have "norm (content K *\<^sub>R f x) \<le> (real n + 1) * (content K * indicator S x)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2147 | proof (cases "x \<in> S") | 
| 53495 | 2148 | case False | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2149 | then show ?thesis by (simp add: 0) | 
| 53495 | 2150 | next | 
| 2151 | case True | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2152 | have *: "content K \<ge> 0" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2153 | using tagged_division_ofD(4)[OF that(1) xk] by auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2154 | moreover have "content K * norm (f x) \<le> content K * (real n + 1)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2155 | by (simp add: mult_left_mono nfx(2)) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2156 | ultimately show ?thesis | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2157 | using nfx True by (auto simp: field_simps) | 
| 53495 | 2158 | qed | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2159 |         ultimately show "\<exists>y. (y, x, K) \<in> (Sigma {..N + 1} q) \<and> norm (content K *\<^sub>R f x) \<le>
 | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2160 | (real y + 1) * (content K *\<^sub>R indicator S x)" | 
| 66199 | 2161 | by force | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2162 | qed auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2163 | also have "... = (\<Sum>i\<le>N + 1. \<Sum>j\<in>q i. (real i + 1) * (case j of (x, K) \<Rightarrow> content K *\<^sub>R indicator S x))" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2164 | apply (rule sum_Sigma_product [symmetric]) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2165 | using q(1) apply auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2166 | done | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2167 | also have "... \<le> (\<Sum>i\<le>N + 1. (real i + 1) * \<bar>\<Sum>(x,K) \<in> q i. content K *\<^sub>R indicator S x\<bar>)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2168 | by (rule sum_mono) (simp add: sum_distrib_left [symmetric]) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2169 | also have "... \<le> (\<Sum>i\<le>N + 1. e/2/2 ^ i)" | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2170 | proof (rule sum_mono) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2171 | show "(real i + 1) * \<bar>\<Sum>(x,K) \<in> q i. content K *\<^sub>R indicator S x\<bar> \<le> e/2/2 ^ i" | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2172 |           if "i \<in> {..N + 1}" for i
 | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2173 | using \<gamma>[of "q i" i] q by (simp add: divide_simps mult.left_commute) | 
| 53495 | 2174 | qed | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2175 | also have "... = e/2 * (\<Sum>i\<le>N + 1. (1/2) ^ i)" | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2176 | unfolding sum_distrib_left by (metis divide_inverse inverse_eq_divide power_one_over) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2177 | also have "\<dots> < e/2 * 2" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2178 | proof (rule mult_strict_left_mono) | 
| 67399 | 2179 |         have "sum ((^) (1/2)) {..N + 1} = sum ((^) (1/2::real)) {..<N + 2}"
 | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2180 | using lessThan_Suc_atMost by auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2181 | also have "... < 2" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2182 | by (auto simp: geometric_sum) | 
| 67399 | 2183 |         finally show "sum ((^) (1/2::real)) {..N + 1} < 2" .
 | 
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2184 | qed (use \<open>0 < e\<close> in auto) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2185 | finally show ?thesis by auto | 
| 53495 | 2186 | qed | 
| 2187 | qed | |
| 2188 | qed | |
| 2189 | ||
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2190 | |
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2191 | proposition has_integral_negligible: | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2192 | fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2193 | assumes negs: "negligible S" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2194 | and "\<And>x. x \<in> (T - S) \<Longrightarrow> f x = 0" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2195 | shows "(f has_integral 0) T" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2196 | proof (cases "\<exists>a b. T = cbox a b") | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2197 | case True | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2198 | then have "((\<lambda>x. if x \<in> T then f x else 0) has_integral 0) T" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2199 | using assms by (auto intro!: has_integral_negligible_cbox) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2200 | then show ?thesis | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2201 | by (rule has_integral_eq [rotated]) auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2202 | next | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2203 | case False | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2204 | let ?f = "(\<lambda>x. if x \<in> T then f x else 0)" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2205 | have "((\<lambda>x. if x \<in> T then f x else 0) has_integral 0) T" | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2206 | apply (auto simp: False has_integral_alt [of ?f]) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2207 | apply (rule_tac x=1 in exI, auto) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2208 | apply (rule_tac x=0 in exI, simp add: has_integral_negligible_cbox [OF negs] assms) | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2209 | done | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2210 | then show ?thesis | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2211 | by (rule_tac f="?f" in has_integral_eq) auto | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2212 | qed | 
| 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2213 | |
| 53495 | 2214 | lemma has_integral_spike: | 
| 56188 | 2215 | fixes f :: "'b::euclidean_space \<Rightarrow> 'a::real_normed_vector" | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2216 | assumes "negligible S" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2217 | and gf: "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2218 | and fint: "(f has_integral y) T" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2219 | shows "(g has_integral y) T" | 
| 53495 | 2220 | proof - | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2221 | have *: "(g has_integral y) (cbox a b)" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2222 | if "(f has_integral y) (cbox a b)" "\<forall>x \<in> cbox a b - S. g x = f x" for a b f and g:: "'b \<Rightarrow> 'a" and y | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2223 | proof - | 
| 56188 | 2224 | have "((\<lambda>x. f x + (g x - f x)) has_integral (y + 0)) (cbox a b)" | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2225 | using that by (intro has_integral_add has_integral_negligible) (auto intro!: \<open>negligible S\<close>) | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2226 | then show ?thesis | 
| 53495 | 2227 | by auto | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2228 | qed | 
| 53495 | 2229 | show ?thesis | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2230 | using fint gf | 
| 53495 | 2231 | apply (subst has_integral_alt) | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2232 | apply (subst (asm) has_integral_alt) | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2233 | apply (simp split: if_split_asm) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2234 | apply (blast dest: *) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2235 | apply (erule_tac V = "\<forall>a b. T \<noteq> cbox a b" in thin_rl) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2236 | apply (elim all_forward imp_forward ex_forward all_forward conj_forward asm_rl) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2237 | apply (auto dest!: *[where f="\<lambda>x. if x\<in>T then f x else 0" and g="\<lambda>x. if x \<in> T then g x else 0"]) | 
| 53495 | 2238 | done | 
| 2239 | qed | |
| 35172 | 2240 | |
| 2241 | lemma has_integral_spike_eq: | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2242 | assumes "negligible S" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2243 | and gf: "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2244 | shows "(f has_integral y) T \<longleftrightarrow> (g has_integral y) T" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2245 | using has_integral_spike [OF \<open>negligible S\<close>] gf | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2246 | by metis | 
| 53495 | 2247 | |
| 2248 | lemma integrable_spike: | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2249 | assumes "negligible S" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2250 | and "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2251 | and "f integrable_on T" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2252 | shows "g integrable_on T" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2253 | using assms unfolding integrable_on_def by (blast intro: has_integral_spike) | 
| 53495 | 2254 | |
| 2255 | lemma integral_spike: | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2256 | assumes "negligible S" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2257 | and "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2258 | shows "integral T f = integral T g" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2259 | using has_integral_spike_eq[OF assms] | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2260 | by (auto simp: integral_def integrable_on_def) | 
| 53495 | 2261 | |
| 35172 | 2262 | |
| 60420 | 2263 | subsection \<open>Some other trivialities about negligible sets.\<close> | 
| 35172 | 2264 | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 2265 | lemma negligible_subset: | 
| 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 2266 | assumes "negligible s" "t \<subseteq> s" | 
| 53495 | 2267 | shows "negligible t" | 
| 2268 | unfolding negligible_def | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 2269 | by (metis (no_types) Diff_iff assms contra_subsetD has_integral_negligible indicator_simps(2)) | 
| 53495 | 2270 | |
| 2271 | lemma negligible_diff[intro?]: | |
| 2272 | assumes "negligible s" | |
| 2273 | shows "negligible (s - t)" | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 2274 | using assms by (meson Diff_subset negligible_subset) | 
| 53495 | 2275 | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2276 | lemma negligible_Int: | 
| 53495 | 2277 | assumes "negligible s \<or> negligible t" | 
| 2278 | shows "negligible (s \<inter> t)" | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 2279 | using assms negligible_subset by force | 
| 53495 | 2280 | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2281 | lemma negligible_Un: | 
| 53495 | 2282 | assumes "negligible s" | 
| 2283 | and "negligible t" | |
| 2284 | shows "negligible (s \<union> t)" | |
| 2285 | unfolding negligible_def | |
| 61166 
5976fe402824
renamed method "goals" to "goal_cases" to emphasize its meaning;
 wenzelm parents: 
61165diff
changeset | 2286 | proof (safe, goal_cases) | 
| 61165 | 2287 | case (1 a b) | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2288 | note assms[unfolded negligible_def,rule_format,of a b] | 
| 53495 | 2289 | then show ?case | 
| 2290 | apply (subst has_integral_spike_eq[OF assms(2)]) | |
| 2291 | defer | |
| 2292 | apply assumption | |
| 2293 | unfolding indicator_def | |
| 2294 | apply auto | |
| 2295 | done | |
| 2296 | qed | |
| 2297 | ||
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2298 | lemma negligible_Un_eq[simp]: "negligible (s \<union> t) \<longleftrightarrow> negligible s \<and> negligible t" | 
| 63956 
b235e845c8e8
HOL-Analysis: add cover lemma ported by L. C. Paulson
 hoelzl parents: 
63945diff
changeset | 2299 | using negligible_Un negligible_subset by blast | 
| 35172 | 2300 | |
| 56188 | 2301 | lemma negligible_sing[intro]: "negligible {a::'a::euclidean_space}"
 | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 2302 | using negligible_standard_hyperplane[OF SOME_Basis, of "a \<bullet> (SOME i. i \<in> Basis)"] negligible_subset by blast | 
| 35172 | 2303 | |
| 53495 | 2304 | lemma negligible_insert[simp]: "negligible (insert a s) \<longleftrightarrow> negligible s" | 
| 2305 | apply (subst insert_is_Un) | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 2306 | unfolding negligible_Un_eq | 
| 53495 | 2307 | apply auto | 
| 2308 | done | |
| 2309 | ||
| 60762 | 2310 | lemma negligible_empty[iff]: "negligible {}"
 | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 2311 | using negligible_insert by blast | 
| 53495 | 2312 | |
| 2313 | lemma negligible_finite[intro]: | |
| 2314 | assumes "finite s" | |
| 2315 | shows "negligible s" | |
| 2316 | using assms by (induct s) auto | |
| 2317 | ||
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 2318 | lemma negligible_Union[intro]: | 
| 66560 | 2319 | assumes "finite \<T>" | 
| 2320 | and "\<And>t. t \<in> \<T> \<Longrightarrow> negligible t" | |
| 2321 | shows "negligible(\<Union>\<T>)" | |
| 53495 | 2322 | using assms by induct auto | 
| 2323 | ||
| 2324 | lemma negligible: | |
| 56188 | 2325 |   "negligible s \<longleftrightarrow> (\<forall>t::('a::euclidean_space) set. ((indicator s::'a\<Rightarrow>real) has_integral 0) t)"
 | 
| 53495 | 2326 | apply safe | 
| 2327 | defer | |
| 2328 | apply (subst negligible_def) | |
| 46905 | 2329 | proof - | 
| 53495 | 2330 | fix t :: "'a set" | 
| 2331 | assume as: "negligible s" | |
| 2332 | have *: "(\<lambda>x. if x \<in> s \<inter> t then 1 else 0) = (\<lambda>x. if x\<in>t then if x\<in>s then 1 else 0 else 0)" | |
| 46905 | 2333 | by auto | 
| 2334 | show "((indicator s::'a\<Rightarrow>real) has_integral 0) t" | |
| 53495 | 2335 | apply (subst has_integral_alt) | 
| 2336 | apply cases | |
| 2337 | apply (subst if_P,assumption) | |
| 46905 | 2338 | unfolding if_not_P | 
| 53495 | 2339 | apply safe | 
| 2340 | apply (rule as[unfolded negligible_def,rule_format]) | |
| 2341 | apply (rule_tac x=1 in exI) | |
| 2342 | apply safe | |
| 2343 | apply (rule zero_less_one) | |
| 2344 | apply (rule_tac x=0 in exI) | |
| 46905 | 2345 | using negligible_subset[OF as,of "s \<inter> t"] | 
| 2346 | unfolding negligible_def indicator_def [abs_def] | |
| 2347 | unfolding * | |
| 2348 | apply auto | |
| 2349 | done | |
| 2350 | qed auto | |
| 35172 | 2351 | |
| 53495 | 2352 | |
| 60420 | 2353 | subsection \<open>Finite case of the spike theorem is quite commonly needed.\<close> | 
| 35172 | 2354 | |
| 53495 | 2355 | lemma has_integral_spike_finite: | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2356 | assumes "finite S" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2357 | and "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2358 | and "(f has_integral y) T" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2359 | shows "(g has_integral y) T" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2360 | using assms has_integral_spike negligible_finite by blast | 
| 53495 | 2361 | |
| 2362 | lemma has_integral_spike_finite_eq: | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2363 | assumes "finite S" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2364 | and "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2365 | shows "((f has_integral y) T \<longleftrightarrow> (g has_integral y) T)" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2366 | by (metis assms has_integral_spike_finite) | 
| 35172 | 2367 | |
| 2368 | lemma integrable_spike_finite: | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2369 | assumes "finite S" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2370 | and "\<And>x. x \<in> T - S \<Longrightarrow> g x = f x" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2371 | and "f integrable_on T" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2372 | shows "g integrable_on T" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
65578diff
changeset | 2373 | using assms has_integral_spike_finite by blast | 
| 53495 | 2374 | |
| 66708 | 2375 | lemma has_integral_bound_spike_finite: | 
| 2376 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | |
| 2377 | assumes "0 \<le> B" "finite S" | |
| 2378 | and f: "(f has_integral i) (cbox a b)" | |
| 2379 | and leB: "\<And>x. x \<in> cbox a b - S \<Longrightarrow> norm (f x) \<le> B" | |
| 2380 | shows "norm i \<le> B * content (cbox a b)" | |
| 2381 | proof - | |
| 2382 | define g where "g \<equiv> (\<lambda>x. if x \<in> S then 0 else f x)" | |
| 2383 | then have "\<And>x. x \<in> cbox a b - S \<Longrightarrow> norm (g x) \<le> B" | |
| 2384 | using leB by simp | |
| 2385 | moreover have "(g has_integral i) (cbox a b)" | |
| 2386 | using has_integral_spike_finite [OF \<open>finite S\<close> _ f] | |
| 2387 | by (simp add: g_def) | |
| 2388 | ultimately show ?thesis | |
| 2389 | by (simp add: \<open>0 \<le> B\<close> g_def has_integral_bound) | |
| 2390 | qed | |
| 2391 | ||
| 2392 | corollary has_integral_bound_real: | |
| 2393 | fixes f :: "real \<Rightarrow> 'b::real_normed_vector" | |
| 2394 | assumes "0 \<le> B" "finite S" | |
| 2395 |       and "(f has_integral i) {a..b}"
 | |
| 2396 |       and "\<And>x. x \<in> {a..b} - S \<Longrightarrow> norm (f x) \<le> B"
 | |
| 2397 |     shows "norm i \<le> B * content {a..b}"
 | |
| 2398 | by (metis assms box_real(2) has_integral_bound_spike_finite) | |
| 2399 | ||
| 35172 | 2400 | |
| 60420 | 2401 | subsection \<open>In particular, the boundary of an interval is negligible.\<close> | 
| 35172 | 2402 | |
| 56188 | 2403 | lemma negligible_frontier_interval: "negligible(cbox (a::'a::euclidean_space) b - box a b)" | 
| 53495 | 2404 | proof - | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 2405 |   let ?A = "\<Union>((\<lambda>k. {x. x\<bullet>k = a\<bullet>k} \<union> {x::'a. x\<bullet>k = b\<bullet>k}) ` Basis)"
 | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2406 | have "negligible ?A" | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2407 | by (force simp add: negligible_Union[OF finite_imageI]) | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2408 | moreover have "cbox a b - box a b \<subseteq> ?A" | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2409 | by (force simp add: mem_box) | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2410 | ultimately show ?thesis | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2411 | by (rule negligible_subset) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 2412 | qed | 
| 35172 | 2413 | |
| 2414 | lemma has_integral_spike_interior: | |
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2415 | assumes f: "(f has_integral y) (cbox a b)" and gf: "\<And>x. x \<in> box a b \<Longrightarrow> g x = f x" | 
| 56188 | 2416 | shows "(g has_integral y) (cbox a b)" | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2417 | apply (rule has_integral_spike[OF negligible_frontier_interval _ f]) | 
| 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2418 | using gf by auto | 
| 35172 | 2419 | |
| 2420 | lemma has_integral_spike_interior_eq: | |
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2421 | assumes "\<And>x. x \<in> box a b \<Longrightarrow> g x = f x" | 
| 56188 | 2422 | shows "(f has_integral y) (cbox a b) \<longleftrightarrow> (g has_integral y) (cbox a b)" | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2423 | by (metis assms has_integral_spike_interior) | 
| 53495 | 2424 | |
| 2425 | lemma integrable_spike_interior: | |
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2426 | assumes "\<And>x. x \<in> box a b \<Longrightarrow> g x = f x" | 
| 56188 | 2427 | and "f integrable_on cbox a b" | 
| 2428 | shows "g integrable_on cbox a b" | |
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 2429 | using assms has_integral_spike_interior_eq by blast | 
| 53495 | 2430 | |
| 35172 | 2431 | |
| 60420 | 2432 | subsection \<open>Integrability of continuous functions.\<close> | 
| 35172 | 2433 | |
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2434 | lemma operative_approximableI: | 
| 61165 | 2435 | fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach" | 
| 53520 | 2436 | assumes "0 \<le> e" | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2437 | shows "operative conj True (\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::'b)) \<le> e) \<and> g integrable_on i)" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2438 | proof - | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2439 | interpret comm_monoid conj True | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2440 | by standard auto | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2441 | show ?thesis | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2442 | proof (standard, safe) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2443 | fix a b :: 'b | 
| 66365 
d77a4ab4fe59
more Henstock_Kurzweil_Integration cleanup
 paulson <lp15@cam.ac.uk> parents: 
66359diff
changeset | 2444 | show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b" | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2445 |       if "box a b = {}" for a b
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2446 | apply (rule_tac x=f in exI) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2447 | using assms that by (auto simp: content_eq_0_interior) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2448 |     {
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2449 | fix c g and k :: 'b | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2450 | assume fg: "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" and g: "g integrable_on cbox a b" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2451 | assume k: "k \<in> Basis" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2452 |       show "\<exists>g. (\<forall>x\<in>cbox a b \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}"
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2453 |            "\<exists>g. (\<forall>x\<in>cbox a b \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}"
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2454 | apply (rule_tac[!] x=g in exI) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2455 | using fg integrable_split[OF g k] by auto | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2456 | } | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2457 | show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2458 |       if fg1: "\<forall>x\<in>cbox a b \<inter> {x. x \<bullet> k \<le> c}. norm (f x - g1 x) \<le> e" 
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2459 |         and g1: "g1 integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}"
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2460 |         and fg2: "\<forall>x\<in>cbox a b \<inter> {x. c \<le> x \<bullet> k}. norm (f x - g2 x) \<le> e" 
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2461 |         and g2: "g2 integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}" 
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2462 | and k: "k \<in> Basis" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2463 | for c k g1 g2 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2464 | proof - | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2465 | let ?g = "\<lambda>x. if x\<bullet>k = c then f x else if x\<bullet>k \<le> c then g1 x else g2 x" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2466 | show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2467 | proof (intro exI conjI ballI) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2468 | show "norm (f x - ?g x) \<le> e" if "x \<in> cbox a b" for x | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2469 | by (auto simp: that assms fg1 fg2) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2470 | show "?g integrable_on cbox a b" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2471 | proof - | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2472 |           have "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c}" "?g integrable_on cbox a b \<inter> {x. x \<bullet> k \<ge> c}"
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2473 | by(rule integrable_spike[OF negligible_standard_hyperplane[of k c]], use k g1 g2 in auto)+ | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2474 | with has_integral_split[OF _ _ k] show ?thesis | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2475 | unfolding integrable_on_def by blast | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2476 | qed | 
| 66365 
d77a4ab4fe59
more Henstock_Kurzweil_Integration cleanup
 paulson <lp15@cam.ac.uk> parents: 
66359diff
changeset | 2477 | qed | 
| 
d77a4ab4fe59
more Henstock_Kurzweil_Integration cleanup
 paulson <lp15@cam.ac.uk> parents: 
66359diff
changeset | 2478 | qed | 
| 53520 | 2479 | qed | 
| 2480 | qed | |
| 2481 | ||
| 67399 | 2482 | lemma comm_monoid_set_F_and: "comm_monoid_set.F (\<and>) True f s \<longleftrightarrow> (finite s \<longrightarrow> (\<forall>x\<in>s. f x))" | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2483 | proof - | 
| 67399 | 2484 | interpret bool: comm_monoid_set "(\<and>)" True | 
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2485 | proof qed auto | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2486 | show ?thesis | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2487 | by (induction s rule: infinite_finite_induct) auto | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2488 | qed | 
| 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 2489 | |
| 53520 | 2490 | lemma approximable_on_division: | 
| 56188 | 2491 | fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach" | 
| 53520 | 2492 | assumes "0 \<le> e" | 
| 66365 
d77a4ab4fe59
more Henstock_Kurzweil_Integration cleanup
 paulson <lp15@cam.ac.uk> parents: 
66359diff
changeset | 2493 | and d: "d division_of (cbox a b)" | 
| 
d77a4ab4fe59
more Henstock_Kurzweil_Integration cleanup
 paulson <lp15@cam.ac.uk> parents: 
66359diff
changeset | 2494 | and f: "\<forall>i\<in>d. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i" | 
| 56188 | 2495 | obtains g where "\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e" "g integrable_on cbox a b" | 
| 53520 | 2496 | proof - | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2497 | interpret operative conj True "\<lambda>i. \<exists>g. (\<forall>x\<in>i. norm (f x - g (x::'b)) \<le> e) \<and> g integrable_on i" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2498 | using \<open>0 \<le> e\<close> by (rule operative_approximableI) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2499 | from f local.division [OF d] that show thesis | 
| 66365 
d77a4ab4fe59
more Henstock_Kurzweil_Integration cleanup
 paulson <lp15@cam.ac.uk> parents: 
66359diff
changeset | 2500 | by auto | 
| 53520 | 2501 | qed | 
| 2502 | ||
| 2503 | lemma integrable_continuous: | |
| 56188 | 2504 | fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach" | 
| 2505 | assumes "continuous_on (cbox a b) f" | |
| 2506 | shows "f integrable_on cbox a b" | |
| 66294 
0442b3f45556
refactored some HORRIBLE integration proofs
 paulson <lp15@cam.ac.uk> parents: 
66199diff
changeset | 2507 | proof (rule integrable_uniform_limit) | 
| 53520 | 2508 | fix e :: real | 
| 2509 | assume e: "e > 0" | |
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2510 | then obtain d where "0 < d" and d: "\<And>x x'. \<lbrakk>x \<in> cbox a b; x' \<in> cbox a b; dist x' x < d\<rbrakk> \<Longrightarrow> dist (f x') (f x) < e" | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2511 | using compact_uniformly_continuous[OF assms compact_cbox] unfolding uniformly_continuous_on_def by metis | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2512 | obtain p where ptag: "p tagged_division_of cbox a b" and finep: "(\<lambda>x. ball x d) fine p" | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2513 | using fine_division_exists[OF gauge_ball[OF \<open>0 < d\<close>], of a b] . | 
| 53520 | 2514 | have *: "\<forall>i\<in>snd ` p. \<exists>g. (\<forall>x\<in>i. norm (f x - g x) \<le> e) \<and> g integrable_on i" | 
| 2515 | proof (safe, unfold snd_conv) | |
| 2516 | fix x l | |
| 2517 | assume as: "(x, l) \<in> p" | |
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2518 | obtain a b where l: "l = cbox a b" | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2519 | using as ptag by blast | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2520 | then have x: "x \<in> cbox a b" | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2521 | using as ptag by auto | 
| 53520 | 2522 | show "\<exists>g. (\<forall>x\<in>l. norm (f x - g x) \<le> e) \<and> g integrable_on l" | 
| 2523 | apply (rule_tac x="\<lambda>y. f x" in exI) | |
| 2524 | proof safe | |
| 2525 | show "(\<lambda>y. f x) integrable_on l" | |
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2526 | unfolding integrable_on_def l by blast | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2527 | next | 
| 53520 | 2528 | fix y | 
| 2529 | assume y: "y \<in> l" | |
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2530 | then have "y \<in> ball x d" | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2531 | using as finep by fastforce | 
| 53520 | 2532 | then show "norm (f y - f x) \<le> e" | 
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2533 | using d x y as l | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2534 | by (metis dist_commute dist_norm less_imp_le mem_ball ptag subsetCE tagged_division_ofD(3)) | 
| 53520 | 2535 | qed | 
| 2536 | qed | |
| 2537 | from e have "e \<ge> 0" | |
| 2538 | by auto | |
| 65578 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2539 | from approximable_on_division[OF this division_of_tagged_division[OF ptag] *] | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2540 | show "\<exists>g. (\<forall>x\<in>cbox a b. norm (f x - g x) \<le> e) \<and> g integrable_on cbox a b" | 
| 
e4997c181cce
New material from PNT proof, as well as more default [simp] declarations. Also removed duplicate theorems about geometric series
 paulson <lp15@cam.ac.uk> parents: 
65204diff
changeset | 2541 | by metis | 
| 53520 | 2542 | qed | 
| 2543 | ||
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 2544 | lemma integrable_continuous_interval: | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 2545 | fixes f :: "'b::ordered_euclidean_space \<Rightarrow> 'a::banach" | 
| 66402 | 2546 |   assumes "continuous_on {a..b} f"
 | 
| 2547 |   shows "f integrable_on {a..b}"
 | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 2548 | by (metis assms integrable_continuous interval_cbox) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 2549 | |
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 2550 | lemmas integrable_continuous_real = integrable_continuous_interval[where 'b=real] | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 2551 | |
| 56188 | 2552 | |
| 60420 | 2553 | subsection \<open>Specialization of additivity to one dimension.\<close> | 
| 35172 | 2554 | |
| 2555 | ||
| 60420 | 2556 | subsection \<open>A useful lemma allowing us to factor out the content size.\<close> | 
| 35172 | 2557 | |
| 2558 | lemma has_integral_factor_content: | |
| 56188 | 2559 | "(f has_integral i) (cbox a b) \<longleftrightarrow> | 
| 2560 | (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> | |
| 64267 | 2561 | norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content (cbox a b)))" | 
| 56188 | 2562 | proof (cases "content (cbox a b) = 0") | 
| 53520 | 2563 | case True | 
| 2564 | show ?thesis | |
| 2565 | unfolding has_integral_null_eq[OF True] | |
| 2566 | apply safe | |
| 2567 | apply (rule, rule, rule gauge_trivial, safe) | |
| 64267 | 2568 | unfolding sum_content_null[OF True] True | 
| 53520 | 2569 | defer | 
| 2570 | apply (erule_tac x=1 in allE) | |
| 2571 | apply safe | |
| 2572 | defer | |
| 2573 | apply (rule fine_division_exists[of _ a b]) | |
| 2574 | apply assumption | |
| 2575 | apply (erule_tac x=p in allE) | |
| 64267 | 2576 | unfolding sum_content_null[OF True] | 
| 53520 | 2577 | apply auto | 
| 2578 | done | |
| 2579 | next | |
| 2580 | case False | |
| 2581 | note F = this[unfolded content_lt_nz[symmetric]] | |
| 2582 | let ?P = "\<lambda>e opp. \<exists>d. gauge d \<and> | |
| 56188 | 2583 | (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> opp (norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i)) e)" | 
| 53520 | 2584 | show ?thesis | 
| 2585 | apply (subst has_integral) | |
| 2586 | proof safe | |
| 2587 | fix e :: real | |
| 2588 | assume e: "e > 0" | |
| 2589 |     {
 | |
| 67399 | 2590 | assume "\<forall>e>0. ?P e (<)" | 
| 2591 | then show "?P (e * content (cbox a b)) (\<le>)" | |
| 56188 | 2592 | apply (erule_tac x="e * content (cbox a b)" in allE) | 
| 53520 | 2593 | apply (erule impE) | 
| 2594 | defer | |
| 2595 | apply (erule exE,rule_tac x=d in exI) | |
| 2596 | using F e | |
| 56544 | 2597 | apply (auto simp add:field_simps) | 
| 53520 | 2598 | done | 
| 2599 | } | |
| 2600 |     {
 | |
| 67399 | 2601 | assume "\<forall>e>0. ?P (e * content (cbox a b)) (\<le>)" | 
| 2602 | then show "?P e (<)" | |
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 2603 | apply (erule_tac x="e/2 / content (cbox a b)" in allE) | 
| 53520 | 2604 | apply (erule impE) | 
| 2605 | defer | |
| 2606 | apply (erule exE,rule_tac x=d in exI) | |
| 2607 | using F e | |
| 56544 | 2608 | apply (auto simp add: field_simps) | 
| 53520 | 2609 | done | 
| 2610 | } | |
| 2611 | qed | |
| 2612 | qed | |
| 2613 | ||
| 56188 | 2614 | lemma has_integral_factor_content_real: | 
| 66402 | 2615 |   "(f has_integral i) {a..b::real} \<longleftrightarrow>
 | 
| 2616 |     (\<forall>e>0. \<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b}  \<and> d fine p \<longrightarrow>
 | |
| 2617 |       norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) p - i) \<le> e * content {a..b} ))"
 | |
| 56188 | 2618 | unfolding box_real[symmetric] | 
| 2619 | by (rule has_integral_factor_content) | |
| 2620 | ||
| 35172 | 2621 | |
| 60420 | 2622 | subsection \<open>Fundamental theorem of calculus.\<close> | 
| 35172 | 2623 | |
| 53520 | 2624 | lemma interval_bounds_real: | 
| 2625 | fixes q b :: real | |
| 2626 | assumes "a \<le> b" | |
| 54777 | 2627 |   shows "Sup {a..b} = b"
 | 
| 2628 |     and "Inf {a..b} = a"
 | |
| 56188 | 2629 | using assms by auto | 
| 53520 | 2630 | |
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2631 | theorem fundamental_theorem_of_calculus: | 
| 53520 | 2632 | fixes f :: "real \<Rightarrow> 'a::banach" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2633 | assumes "a \<le> b" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2634 |       and vecd: "\<And>x. x \<in> {a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x within {a..b})"
 | 
| 66402 | 2635 |   shows "(f' has_integral (f b - f a)) {a..b}"
 | 
| 56188 | 2636 | unfolding has_integral_factor_content box_real[symmetric] | 
| 53520 | 2637 | proof safe | 
| 2638 | fix e :: real | |
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 2639 | assume "e > 0" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2640 |   then have "\<forall>x. \<exists>d>0. x \<in> {a..b} \<longrightarrow>
 | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2641 |          (\<forall>y\<in>{a..b}. norm (y-x) < d \<longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e * norm (y-x))"
 | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 2642 | using vecd[unfolded has_vector_derivative_def has_derivative_within_alt] by blast | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 2643 | then obtain d where d: "\<And>x. 0 < d x" | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 2644 |                  "\<And>x y. \<lbrakk>x \<in> {a..b}; y \<in> {a..b}; norm (y-x) < d x\<rbrakk>
 | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 2645 | \<Longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e * norm (y-x)" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2646 | by metis | 
| 56188 | 2647 | show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of (cbox a b) \<and> d fine p \<longrightarrow> | 
| 2648 | norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b))" | |
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2649 | proof (rule exI, safe) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2650 | show "gauge (\<lambda>x. ball x (d x))" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2651 | using d(1) gauge_ball_dependent by blast | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2652 | next | 
| 53520 | 2653 | fix p | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2654 | assume ptag: "p tagged_division_of cbox a b" and finep: "(\<lambda>x. ball x (d x)) fine p" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2655 | have ba: "b - a = (\<Sum>(x,K)\<in>p. Sup K - Inf K)" "f b - f a = (\<Sum>(x,K)\<in>p. f(Sup K) - f(Inf K))" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2656 | using additive_tagged_division_1[where f= "\<lambda>x. x"] additive_tagged_division_1[where f= f] | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2657 | \<open>a \<le> b\<close> ptag by auto | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2658 | have "norm (\<Sum>(x, K) \<in> p. (content K *\<^sub>R f' x) - (f (Sup K) - f (Inf K))) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2659 | \<le> (\<Sum>n\<in>p. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))" | 
| 64267 | 2660 | proof (rule sum_norm_le,safe) | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2661 | fix x K | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2662 | assume "(x, K) \<in> p" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2663 | then have "x \<in> K" and kab: "K \<subseteq> cbox a b" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2664 | using ptag by blast+ | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2665 | then obtain u v where k: "K = cbox u v" and "x \<in> K" and kab: "K \<subseteq> cbox a b" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2666 | using ptag \<open>(x, K) \<in> p\<close> by auto | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2667 | have "u \<le> v" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2668 | using \<open>x \<in> K\<close> unfolding k by auto | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2669 | have ball: "\<forall>y\<in>K. y \<in> ball x (d x)" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2670 | using finep \<open>(x, K) \<in> p\<close> by blast | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2671 | have "u \<in> K" "v \<in> K" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2672 | by (simp_all add: \<open>u \<le> v\<close> k) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2673 | have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = norm (f u - f x - (u - x) *\<^sub>R f' x - (f v - f x - (v - x) *\<^sub>R f' x))" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2674 | by (auto simp add: algebra_simps) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2675 | also have "... \<le> norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2676 | by (rule norm_triangle_ineq4) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2677 | finally have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) \<le> | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2678 | norm (f u - f x - (u - x) *\<^sub>R f' x) + norm (f v - f x - (v - x) *\<^sub>R f' x)" . | 
| 53520 | 2679 | also have "\<dots> \<le> e * norm (u - x) + e * norm (v - x)" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2680 | proof (rule add_mono) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2681 | show "norm (f u - f x - (u - x) *\<^sub>R f' x) \<le> e * norm (u - x)" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2682 | apply (rule d(2)[of x u]) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2683 | using \<open>x \<in> K\<close> kab \<open>u \<in> K\<close> ball dist_real_def by (auto simp add:dist_real_def) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2684 | show "norm (f v - f x - (v - x) *\<^sub>R f' x) \<le> e * norm (v - x)" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2685 | apply (rule d(2)[of x v]) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2686 | using \<open>x \<in> K\<close> kab \<open>v \<in> K\<close> ball dist_real_def by (auto simp add:dist_real_def) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2687 | qed | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2688 | also have "\<dots> \<le> e * (Sup K - Inf K)" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2689 | using \<open>x \<in> K\<close> by (auto simp: k interval_bounds_real[OF \<open>u \<le> v\<close>] field_simps) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2690 | finally show "norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (Sup K - Inf K)" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2691 | using \<open>u \<le> v\<close> by (simp add: k) | 
| 53520 | 2692 | qed | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2693 | with \<open>a \<le> b\<close> show "norm ((\<Sum>(x, K)\<in>p. content K *\<^sub>R f' x) - (f b - f a)) \<le> e * content (cbox a b)" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2694 | by (auto simp: ba split_def sum_subtractf [symmetric] sum_distrib_left) | 
| 53520 | 2695 | qed | 
| 2696 | qed | |
| 2697 | ||
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2698 | lemma ident_has_integral: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2699 | fixes a::real | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2700 | assumes "a \<le> b" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2701 |   shows "((\<lambda>x. x) has_integral (b\<^sup>2 - a\<^sup>2)/2) {a..b}"
 | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2702 | proof - | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2703 |   have "((\<lambda>x. x) has_integral inverse 2 * b\<^sup>2 - inverse 2 * a\<^sup>2) {a..b}"
 | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 2704 | apply (rule fundamental_theorem_of_calculus [OF assms]) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2705 | unfolding power2_eq_square | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2706 | by (rule derivative_eq_intros | simp)+ | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2707 | then show ?thesis | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2708 | by (simp add: field_simps) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2709 | qed | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2710 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2711 | lemma integral_ident [simp]: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2712 | fixes a::real | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2713 | assumes "a \<le> b" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2714 |   shows "integral {a..b} (\<lambda>x. x) = (if a \<le> b then (b\<^sup>2 - a\<^sup>2)/2 else 0)"
 | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2715 | by (metis assms ident_has_integral integral_unique) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2716 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2717 | lemma ident_integrable_on: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2718 | fixes a::real | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2719 |   shows "(\<lambda>x. x) integrable_on {a..b}"
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2720 | by (metis atLeastatMost_empty_iff integrable_on_def has_integral_empty ident_has_integral) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 2721 | |
| 35172 | 2722 | |
| 60420 | 2723 | subsection \<open>Taylor series expansion\<close> | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2724 | |
| 64267 | 2725 | lemma (in bounded_bilinear) sum_prod_derivatives_has_vector_derivative: | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2726 | assumes "p>0" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2727 | and f0: "Df 0 = f" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2728 | and Df: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow> | 
| 66402 | 2729 |     (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})"
 | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2730 | and g0: "Dg 0 = g" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2731 | and Dg: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow> | 
| 66402 | 2732 |     (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})"
 | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2733 | and ivl: "a \<le> t" "t \<le> b" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2734 | shows "((\<lambda>t. \<Sum>i<p. (-1)^i *\<^sub>R prod (Df i t) (Dg (p - Suc i) t)) | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2735 | has_vector_derivative | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2736 | prod (f t) (Dg p t) - (-1)^p *\<^sub>R prod (Df p t) (g t)) | 
| 66402 | 2737 |     (at t within {a..b})"
 | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2738 | using assms | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2739 | proof cases | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2740 | assume p: "p \<noteq> 1" | 
| 63040 | 2741 | define p' where "p' = p - 2" | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2742 |   from assms p have p': "{..<p} = {..Suc p'}" "p = Suc (Suc p')"
 | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2743 | by (auto simp: p'_def) | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2744 | have *: "\<And>i. i \<le> p' \<Longrightarrow> Suc (Suc p' - i) = (Suc (Suc p') - i)" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2745 | by auto | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2746 | let ?f = "\<lambda>i. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg ((p - i)) t))" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2747 | have "(\<Sum>i<p. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg (Suc (p - Suc i)) t) + | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2748 | prod (Df (Suc i) t) (Dg (p - Suc i) t))) = | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2749 | (\<Sum>i\<le>(Suc p'). ?f i - ?f (Suc i))" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2750 | by (auto simp: algebra_simps p'(2) numeral_2_eq_2 * lessThan_Suc_atMost) | 
| 64267 | 2751 | also note sum_telescope | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2752 | finally | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2753 | have "(\<Sum>i<p. (-1) ^ i *\<^sub>R (prod (Df i t) (Dg (Suc (p - Suc i)) t) + | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2754 | prod (Df (Suc i) t) (Dg (p - Suc i) t))) | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2755 | = prod (f t) (Dg p t) - (- 1) ^ p *\<^sub>R prod (Df p t) (g t)" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2756 | unfolding p'[symmetric] | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2757 | by (simp add: assms) | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2758 | thus ?thesis | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2759 | using assms | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2760 | by (auto intro!: derivative_eq_intros has_vector_derivative) | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2761 | qed (auto intro!: derivative_eq_intros has_vector_derivative) | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2762 | |
| 60621 | 2763 | lemma | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2764 | fixes f::"real\<Rightarrow>'a::banach" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2765 | assumes "p>0" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2766 | and f0: "Df 0 = f" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2767 | and Df: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow> | 
| 66402 | 2768 |     (Df m has_vector_derivative Df (Suc m) t) (at t within {a..b})"
 | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2769 | and ivl: "a \<le> b" | 
| 60621 | 2770 | defines "i \<equiv> \<lambda>x. ((b - x) ^ (p - 1) / fact (p - 1)) *\<^sub>R Df p x" | 
| 2771 | shows taylor_has_integral: | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2772 |     "(i has_integral f b - (\<Sum>i<p. ((b-a) ^ i / fact i) *\<^sub>R Df i a)) {a..b}"
 | 
| 60621 | 2773 | and taylor_integral: | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2774 |     "f b = (\<Sum>i<p. ((b-a) ^ i / fact i) *\<^sub>R Df i a) + integral {a..b} i"
 | 
| 60621 | 2775 | and taylor_integrable: | 
| 66402 | 2776 |     "i integrable_on {a..b}"
 | 
| 61166 
5976fe402824
renamed method "goals" to "goal_cases" to emphasize its meaning;
 wenzelm parents: 
61165diff
changeset | 2777 | proof goal_cases | 
| 60621 | 2778 | case 1 | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2779 | interpret bounded_bilinear "scaleR::real\<Rightarrow>'a\<Rightarrow>'a" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2780 | by (rule bounded_bilinear_scaleR) | 
| 63040 | 2781 | define g where "g s = (b - s)^(p - 1)/fact (p - 1)" for s | 
| 2782 | define Dg where [abs_def]: | |
| 2783 | "Dg n s = (if n < p then (-1)^n * (b - s)^(p - 1 - n) / fact (p - 1 - n) else 0)" for n s | |
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2784 | have g0: "Dg 0 = g" | 
| 60420 | 2785 | using \<open>p > 0\<close> | 
| 62390 | 2786 | by (auto simp add: Dg_def divide_simps g_def split: if_split_asm) | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2787 |   {
 | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2788 | fix m | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2789 | assume "p > Suc m" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2790 | hence "p - Suc m = Suc (p - Suc (Suc m))" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2791 | by auto | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2792 | hence "real (p - Suc m) * fact (p - Suc (Suc m)) = fact (p - Suc m)" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2793 | by auto | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2794 | } note fact_eq = this | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2795 | have Dg: "\<And>m t. m < p \<Longrightarrow> a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow> | 
| 66402 | 2796 |     (Dg m has_vector_derivative Dg (Suc m) t) (at t within {a..b})"
 | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2797 | unfolding Dg_def | 
| 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: 
61524diff
changeset | 2798 | by (auto intro!: derivative_eq_intros simp: has_vector_derivative_def fact_eq divide_simps) | 
| 60621 | 2799 | let ?sum = "\<lambda>t. \<Sum>i<p. (- 1) ^ i *\<^sub>R Dg i t *\<^sub>R Df (p - Suc i) t" | 
| 64267 | 2800 | from sum_prod_derivatives_has_vector_derivative[of _ Dg _ _ _ Df, | 
| 60420 | 2801 | OF \<open>p > 0\<close> g0 Dg f0 Df] | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2802 | have deriv: "\<And>t. a \<le> t \<Longrightarrow> t \<le> b \<Longrightarrow> | 
| 60621 | 2803 | (?sum has_vector_derivative | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2804 |       g t *\<^sub>R Df p t - (- 1) ^ p *\<^sub>R Dg p t *\<^sub>R f t) (at t within {a..b})"
 | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2805 | by auto | 
| 60420 | 2806 | from fundamental_theorem_of_calculus[rule_format, OF \<open>a \<le> b\<close> deriv] | 
| 66402 | 2807 |   have "(i has_integral ?sum b - ?sum a) {a..b}"
 | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 2808 | using atLeastatMost_empty'[simp del] | 
| 60621 | 2809 | by (simp add: i_def g_def Dg_def) | 
| 2810 | also | |
| 2811 | have one: "(- 1) ^ p' * (- 1) ^ p' = (1::real)" | |
| 2812 |     and "{..<p} \<inter> {i. p = Suc i} = {p - 1}"
 | |
| 2813 | for p' | |
| 61222 | 2814 | using \<open>p > 0\<close> | 
| 60621 | 2815 | by (auto simp: power_mult_distrib[symmetric]) | 
| 2816 | then have "?sum b = f b" | |
| 61222 | 2817 | using Suc_pred'[OF \<open>p > 0\<close>] | 
| 60621 | 2818 | by (simp add: diff_eq_eq Dg_def power_0_left le_Suc_eq if_distrib | 
| 64267 | 2819 | cond_application_beta sum.If_cases f0) | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2820 | also | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2821 |   have "{..<p} = (\<lambda>x. p - x - 1) ` {..<p}"
 | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2822 | proof safe | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2823 | fix x | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2824 | assume "x < p" | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2825 |     thus "x \<in> (\<lambda>x. p - x - 1) ` {..<p}"
 | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2826 | by (auto intro!: image_eqI[where x = "p - x - 1"]) | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2827 | qed simp | 
| 60621 | 2828 | from _ this | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2829 | have "?sum a = (\<Sum>i<p. ((b-a) ^ i / fact i) *\<^sub>R Df i a)" | 
| 64267 | 2830 | by (rule sum.reindex_cong) (auto simp add: inj_on_def Dg_def one) | 
| 60621 | 2831 | finally show c: ?case . | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2832 | case 2 show ?case using c integral_unique | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 2833 | by (metis (lifting) add.commute diff_eq_eq integral_unique) | 
| 60621 | 2834 | case 3 show ?case using c by force | 
| 60180 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2835 | qed | 
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2836 | |
| 
09a7481c03b1
general Taylor series expansion with integral remainder
 immler parents: 
59765diff
changeset | 2837 | |
| 35172 | 2838 | |
| 60420 | 2839 | subsection \<open>Only need trivial subintervals if the interval itself is trivial.\<close> | 
| 35172 | 2840 | |
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2841 | proposition division_of_nontrivial: | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2842 | fixes \<D> :: "'a::euclidean_space set set" | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2843 | assumes sdiv: "\<D> division_of (cbox a b)" | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2844 | and cont0: "content (cbox a b) \<noteq> 0" | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2845 |   shows "{k. k \<in> \<D> \<and> content k \<noteq> 0} division_of (cbox a b)"
 | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2846 | using sdiv | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2847 | proof (induction "card \<D>" arbitrary: \<D> rule: less_induct) | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2848 | case less | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2849 | note \<D> = division_ofD[OF less.prems] | 
| 53520 | 2850 |   {
 | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2851 |     presume *: "{k \<in> \<D>. content k \<noteq> 0} \<noteq> \<D> \<Longrightarrow> ?case"
 | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2852 | then show ?case | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2853 | using less.prems by fastforce | 
| 53520 | 2854 | } | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2855 |   assume noteq: "{k \<in> \<D>. content k \<noteq> 0} \<noteq> \<D>"
 | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2856 | then obtain K c d where "K \<in> \<D>" and contk: "content K = 0" and keq: "K = cbox c d" | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2857 | using \<D>(4) by blast | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2858 | then have "card \<D> > 0" | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2859 | unfolding card_gt_0_iff using less by auto | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2860 |   then have card: "card (\<D> - {K}) < card \<D>"
 | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2861 | using less \<open>K \<in> \<D>\<close> by (simp add: \<D>(1)) | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2862 |   have closed: "closed (\<Union>(\<D> - {K}))"
 | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2863 | using less.prems by auto | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2864 |   have "x islimpt \<Union>(\<D> - {K})" if "x \<in> K" for x 
 | 
| 53520 | 2865 | unfolding islimpt_approachable | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2866 | proof (intro allI impI) | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2867 | fix e::real | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2868 | assume "e > 0" | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 2869 | obtain i where i: "c\<bullet>i = d\<bullet>i" "i\<in>Basis" | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2870 | using contk \<D>(3) [OF \<open>K \<in> \<D>\<close>] unfolding box_ne_empty keq | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2871 | by (meson content_eq_0 dual_order.antisym) | 
| 53520 | 2872 | then have xi: "x\<bullet>i = d\<bullet>i" | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2873 | using \<open>x \<in> K\<close> unfolding keq mem_box by (metis antisym) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2874 | define y where "y = (\<Sum>j\<in>Basis. (if j = i then if c\<bullet>i \<le> (a\<bullet>i + b\<bullet>i)/2 then c\<bullet>i + | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 2875 | min e (b\<bullet>i - c\<bullet>i)/2 else c\<bullet>i - min e (c\<bullet>i - a\<bullet>i)/2 else x\<bullet>j) *\<^sub>R j)" | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2876 |     show "\<exists>x'\<in>\<Union>(\<D> - {K}). x' \<noteq> x \<and> dist x' x < e"
 | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2877 | proof (intro bexI conjI) | 
| 56188 | 2878 | have "d \<in> cbox c d" | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2879 | using \<D>(3)[OF \<open>K \<in> \<D>\<close>] by (simp add: box_ne_empty(1) keq mem_box(2)) | 
| 56188 | 2880 | then have "d \<in> cbox a b" | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2881 | using \<D>(2)[OF \<open>K \<in> \<D>\<close>] by (auto simp: keq) | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2882 | then have di: "a \<bullet> i \<le> d \<bullet> i \<and> d \<bullet> i \<le> b \<bullet> i" | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2883 | using \<open>i \<in> Basis\<close> mem_box(2) by blast | 
| 53520 | 2884 | then have xyi: "y\<bullet>i \<noteq> x\<bullet>i" | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2885 | unfolding y_def i xi using \<open>e > 0\<close> cont0 \<open>i \<in> Basis\<close> | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2886 | by (auto simp: content_eq_0 elim!: ballE[of _ _ i]) | 
| 53520 | 2887 | then show "y \<noteq> x" | 
| 2888 | unfolding euclidean_eq_iff[where 'a='a] using i by auto | |
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2889 | have "norm (y-x) \<le> (\<Sum>b\<in>Basis. \<bar>(y - x) \<bullet> b\<bar>)" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2890 | by (rule norm_le_l1) | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2891 |       also have "... = \<bar>(y - x) \<bullet> i\<bar> + (\<Sum>b \<in> Basis - {i}. \<bar>(y - x) \<bullet> b\<bar>)"
 | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2892 | by (meson finite_Basis i(2) sum.remove) | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2893 | also have "... < e + sum (\<lambda>i. 0) Basis" | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2894 | proof (rule add_less_le_mono) | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 2895 | show "\<bar>(y-x) \<bullet> i\<bar> < e" | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2896 | using di \<open>e > 0\<close> y_def i xi by (auto simp: inner_simps) | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 2897 |         show "(\<Sum>i\<in>Basis - {i}. \<bar>(y-x) \<bullet> i\<bar>) \<le> (\<Sum>i\<in>Basis. 0)"
 | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 2898 | unfolding y_def by (auto simp: inner_simps) | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2899 | qed | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2900 | finally have "norm (y-x) < e + sum (\<lambda>i. 0) Basis" . | 
| 53520 | 2901 | then show "dist y x < e" | 
| 2902 | unfolding dist_norm by auto | |
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2903 | have "y \<notin> K" | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2904 | unfolding keq mem_box using i(1) i(2) xi xyi by fastforce | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2905 | moreover have "y \<in> \<Union>\<D>" | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2906 | using subsetD[OF \<D>(2)[OF \<open>K \<in> \<D>\<close>] \<open>x \<in> K\<close>] \<open>e > 0\<close> di i | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2907 | by (auto simp: \<D> mem_box y_def field_simps elim!: ballE[of _ _ i]) | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2908 |       ultimately show "y \<in> \<Union>(\<D> - {K})" by auto
 | 
| 53520 | 2909 | qed | 
| 2910 | qed | |
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2911 |   then have "K \<subseteq> \<Union>(\<D> - {K})"
 | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2912 | using closed closed_limpt by blast | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2913 |   then have "\<Union>(\<D> - {K}) = cbox a b"
 | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2914 | unfolding \<D>(6)[symmetric] by auto | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2915 |   then have "\<D> - {K} division_of cbox a b"
 | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2916 | by (metis Diff_subset less.prems division_of_subset \<D>(6)) | 
| 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2917 |   then have "{ka \<in> \<D> - {K}. content ka \<noteq> 0} division_of (cbox a b)"
 | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2918 | using card less.hyps by blast | 
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2919 |   moreover have "{ka \<in> \<D> - {K}. content ka \<noteq> 0} = {K \<in> \<D>. content K \<noteq> 0}"
 | 
| 66523 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2920 | using contk by auto | 
| 
5a1a2ac950c2
division_of_nontrivial partial cleanup
 paulson <lp15@cam.ac.uk> parents: 
66519diff
changeset | 2921 | ultimately show ?case by auto | 
| 53520 | 2922 | qed | 
| 2923 | ||
| 35172 | 2924 | |
| 60420 | 2925 | subsection \<open>Integrability on subintervals.\<close> | 
| 35172 | 2926 | |
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2927 | lemma operative_integrableI: | 
| 56188 | 2928 | fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach" | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2929 | assumes "0 \<le> e" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2930 | shows "operative conj True (\<lambda>i. f integrable_on i)" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2931 | proof - | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2932 | interpret comm_monoid conj True | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2933 | by standard auto | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2934 |   have 1: "\<And>a b. box a b = {} \<Longrightarrow> f integrable_on cbox a b"
 | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2935 | by (simp add: content_eq_0_interior integrable_on_null) | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2936 | have 2: "\<And>a b c k. | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2937 | \<lbrakk>k \<in> Basis; | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2938 |         f integrable_on cbox a b \<inter> {x. x \<bullet> k \<le> c};
 | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2939 |         f integrable_on cbox a b \<inter> {x. c \<le> x \<bullet> k}\<rbrakk>
 | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2940 | \<Longrightarrow> f integrable_on cbox a b" | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2941 | unfolding integrable_on_def by (auto intro!: has_integral_split) | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2942 | show ?thesis | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2943 | apply standard | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 2944 | using 1 2 by blast+ | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2945 | qed | 
| 53520 | 2946 | |
| 2947 | lemma integrable_subinterval: | |
| 56188 | 2948 | fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach" | 
| 2949 | assumes "f integrable_on cbox a b" | |
| 2950 | and "cbox c d \<subseteq> cbox a b" | |
| 2951 | shows "f integrable_on cbox c d" | |
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2952 | proof - | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2953 | interpret operative conj True "\<lambda>i. f integrable_on i" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2954 | using order_refl by (rule operative_integrableI) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2955 | show ?thesis | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2956 |     apply (cases "cbox c d = {}")
 | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2957 | defer | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2958 | apply (rule partial_division_extend_1[OF assms(2)],assumption) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2959 | using division [symmetric] assms(1) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2960 | apply (auto simp: comm_monoid_set_F_and) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2961 | done | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2962 | qed | 
| 53520 | 2963 | |
| 56188 | 2964 | lemma integrable_subinterval_real: | 
| 2965 | fixes f :: "real \<Rightarrow> 'a::banach" | |
| 66402 | 2966 |   assumes "f integrable_on {a..b}"
 | 
| 2967 |     and "{c..d} \<subseteq> {a..b}"
 | |
| 2968 |   shows "f integrable_on {c..d}"
 | |
| 66524 
0d8dab1f6903
some tidying of division_of_nontrivial
 paulson <lp15@cam.ac.uk> parents: 
66523diff
changeset | 2969 | by (metis assms box_real(2) integrable_subinterval) | 
| 56188 | 2970 | |
| 35172 | 2971 | |
| 60420 | 2972 | subsection \<open>Combining adjacent intervals in 1 dimension.\<close> | 
| 35172 | 2973 | |
| 53520 | 2974 | lemma has_integral_combine: | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2975 | fixes a b c :: real and j :: "'a::banach" | 
| 53520 | 2976 | assumes "a \<le> c" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2977 | and "c \<le> b" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2978 |       and ac: "(f has_integral i) {a..c}"
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2979 |       and cb: "(f has_integral j) {c..b}"
 | 
| 66402 | 2980 |   shows "(f has_integral (i + j)) {a..b}"
 | 
| 53520 | 2981 | proof - | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2982 | interpret operative_real "lift_option plus" "Some 0" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2983 | "\<lambda>i. if f integrable_on i then Some (integral i f) else None" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2984 | using operative_integralI by (rule operative_realI) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 2985 | from \<open>a \<le> c\<close> \<open>c \<le> b\<close> ac cb coalesce_less_eq | 
| 67399 | 2986 | have *: "lift_option (+) | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2987 |              (if f integrable_on {a..c} then Some (integral {a..c} f) else None)
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2988 |              (if f integrable_on {c..b} then Some (integral {c..b} f) else None) =
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2989 |             (if f integrable_on {a..b} then Some (integral {a..b} f) else None)"
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2990 | by (auto simp: split: if_split_asm) | 
| 56188 | 2991 | then have "f integrable_on cbox a b" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2992 | using ac cb by (auto split: if_split_asm) | 
| 53520 | 2993 | with * | 
| 2994 | show ?thesis | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 2995 | using ac cb by (auto simp add: integrable_on_def integral_unique split: if_split_asm) | 
| 53520 | 2996 | qed | 
| 2997 | ||
| 2998 | lemma integral_combine: | |
| 2999 | fixes f :: "real \<Rightarrow> 'a::banach" | |
| 3000 | assumes "a \<le> c" | |
| 3001 | and "c \<le> b" | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3002 |     and ab: "f integrable_on {a..b}"
 | 
| 66402 | 3003 |   shows "integral {a..c} f + integral {c..b} f = integral {a..b} f"
 | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3004 | proof - | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3005 |   have "(f has_integral integral {a..c} f) {a..c}"
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3006 | using ab \<open>c \<le> b\<close> integrable_subinterval_real by fastforce | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3007 | moreover | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3008 |   have "(f has_integral integral {c..b} f) {c..b}"
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3009 | using ab \<open>a \<le> c\<close> integrable_subinterval_real by fastforce | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3010 |   ultimately have "(f has_integral integral {a..c} f + integral {c..b} f) {a..b}"
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3011 | using \<open>a \<le> c\<close> \<open>c \<le> b\<close> has_integral_combine by blast | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3012 | then show ?thesis | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3013 | by (simp add: has_integral_integrable_integral) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3014 | qed | 
| 53520 | 3015 | |
| 3016 | lemma integrable_combine: | |
| 3017 | fixes f :: "real \<Rightarrow> 'a::banach" | |
| 3018 | assumes "a \<le> c" | |
| 3019 | and "c \<le> b" | |
| 66402 | 3020 |     and "f integrable_on {a..c}"
 | 
| 3021 |     and "f integrable_on {c..b}"
 | |
| 3022 |   shows "f integrable_on {a..b}"
 | |
| 53520 | 3023 | using assms | 
| 3024 | unfolding integrable_on_def | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 3025 | by (auto intro!: has_integral_combine) | 
| 53520 | 3026 | |
| 35172 | 3027 | |
| 60420 | 3028 | subsection \<open>Reduce integrability to "local" integrability.\<close> | 
| 35172 | 3029 | |
| 53520 | 3030 | lemma integrable_on_little_subintervals: | 
| 56188 | 3031 | fixes f :: "'b::euclidean_space \<Rightarrow> 'a::banach" | 
| 3032 | assumes "\<forall>x\<in>cbox a b. \<exists>d>0. \<forall>u v. x \<in> cbox u v \<and> cbox u v \<subseteq> ball x d \<and> cbox u v \<subseteq> cbox a b \<longrightarrow> | |
| 3033 | f integrable_on cbox u v" | |
| 3034 | shows "f integrable_on cbox a b" | |
| 3035 | proof - | |
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 3036 | interpret operative conj True "\<lambda>i. f integrable_on i" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 3037 | using order_refl by (rule operative_integrableI) | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3038 | have "\<forall>x. \<exists>d>0. x\<in>cbox a b \<longrightarrow> (\<forall>u v. x \<in> cbox u v \<and> cbox u v \<subseteq> ball x d \<and> cbox u v \<subseteq> cbox a b \<longrightarrow> | 
| 56188 | 3039 | f integrable_on cbox u v)" | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3040 | using assms by (metis zero_less_one) | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3041 | then obtain d where d: "\<And>x. 0 < d x" | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3042 | "\<And>x u v. \<lbrakk>x \<in> cbox a b; x \<in> cbox u v; cbox u v \<subseteq> ball x (d x); cbox u v \<subseteq> cbox a b\<rbrakk> | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3043 | \<Longrightarrow> f integrable_on cbox u v" | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3044 | by metis | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3045 | obtain p where p: "p tagged_division_of cbox a b" "(\<lambda>x. ball x (d x)) fine p" | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3046 | using fine_division_exists[OF gauge_ball_dependent,of d a b] d(1) by blast | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3047 | then have sndp: "snd ` p division_of cbox a b" | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3048 | by (metis division_of_tagged_division) | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3049 | have "f integrable_on k" if "(x, k) \<in> p" for x k | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3050 | using tagged_division_ofD(2-4)[OF p(1) that] fineD[OF p(2) that] d[of x] by auto | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3051 | then show ?thesis | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 3052 | unfolding division [symmetric, OF sndp] comm_monoid_set_F_and | 
| 53520 | 3053 | by auto | 
| 3054 | qed | |
| 3055 | ||
| 35172 | 3056 | |
| 63593 
bbcb05504fdc
HOL-Multivariate_Analysis: replace neutral, monoidal, and iterate by the comm_monoid_set versions. Changed operative to comm_monoid_set. Renamed support_on to support and changed to comm_monoid_add.
 hoelzl parents: 
63540diff
changeset | 3057 | subsection \<open>Second FTC or existence of antiderivative.\<close> | 
| 35172 | 3058 | |
| 56188 | 3059 | lemma integrable_const[intro]: "(\<lambda>x. c) integrable_on cbox a b" | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3060 | unfolding integrable_on_def by blast | 
| 53520 | 3061 | |
| 61204 | 3062 | lemma integral_has_vector_derivative_continuous_at: | 
| 3063 | fixes f :: "real \<Rightarrow> 'a::banach" | |
| 3064 |   assumes f: "f integrable_on {a..b}"
 | |
| 66708 | 3065 |      and x: "x \<in> {a..b} - S"
 | 
| 3066 | and "finite S" | |
| 3067 |      and fx: "continuous (at x within ({a..b} - S)) f"
 | |
| 3068 |  shows "((\<lambda>u. integral {a..u} f) has_vector_derivative f x) (at x within ({a..b} - S))"
 | |
| 61204 | 3069 | proof - | 
| 3070 |   let ?I = "\<lambda>a b. integral {a..b} f"
 | |
| 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: 
61524diff
changeset | 3071 |   { fix e::real
 | 
| 61204 | 3072 | assume "e > 0" | 
| 66708 | 3073 |     obtain d where "d>0" and d: "\<And>x'. \<lbrakk>x' \<in> {a..b} - S; \<bar>x' - x\<bar> < d\<rbrakk> \<Longrightarrow> norm(f x' - f x) \<le> e"
 | 
| 61222 | 3074 | using \<open>e>0\<close> fx by (auto simp: continuous_within_eps_delta dist_norm less_imp_le) | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3075 |     have "norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
 | 
| 66708 | 3076 |            if y: "y \<in> {a..b} - S" and yx: "\<bar>y - x\<bar> < d" for y
 | 
| 61204 | 3077 | proof (cases "y < x") | 
| 3078 | case False | |
| 3079 |       have "f integrable_on {a..y}"
 | |
| 3080 | using f y by (simp add: integrable_subinterval_real) | |
| 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: 
61524diff
changeset | 3081 | then have Idiff: "?I a y - ?I a x = ?I x y" | 
| 61204 | 3082 | using False x by (simp add: algebra_simps integral_combine) | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3083 |       have fux_int: "((\<lambda>u. f u - f x) has_integral integral {x..y} f - (y-x) *\<^sub>R f x) {x..y}"
 | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 3084 | apply (rule has_integral_diff) | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 3085 | using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) | 
| 61204 | 3086 | using has_integral_const_real [of "f x" x y] False | 
| 66708 | 3087 | apply simp | 
| 61204 | 3088 | done | 
| 66708 | 3089 | have "\<And>xa. y - x < d \<Longrightarrow> (\<And>x'. a \<le> x' \<and> x' \<le> b \<and> x' \<notin> S \<Longrightarrow> \<bar>x' - x\<bar> < d \<Longrightarrow> norm (f x' - f x) \<le> e) \<Longrightarrow> 0 < e \<Longrightarrow> xa \<notin> S \<Longrightarrow> a \<le> x \<Longrightarrow> x \<notin> S \<Longrightarrow> y \<le> b \<Longrightarrow> y \<notin> S \<Longrightarrow> x \<le> xa \<Longrightarrow> xa \<le> y \<Longrightarrow> norm (f xa - f x) \<le> e" | 
| 3090 | using assms by auto | |
| 61204 | 3091 | show ?thesis | 
| 3092 | using False | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3093 | apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) | 
| 61204 | 3094 | apply (rule has_integral_bound_real[where f="(\<lambda>u. f u - f x)"]) | 
| 66708 | 3095 | using yx False d x y \<open>e>0\<close> assms by (auto simp: Idiff fux_int) | 
| 61204 | 3096 | next | 
| 3097 | case True | |
| 3098 |       have "f integrable_on {a..x}"
 | |
| 3099 | using f x by (simp add: integrable_subinterval_real) | |
| 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: 
61524diff
changeset | 3100 | then have Idiff: "?I a x - ?I a y = ?I y x" | 
| 61204 | 3101 | using True x y by (simp add: algebra_simps integral_combine) | 
| 3102 |       have fux_int: "((\<lambda>u. f u - f x) has_integral integral {y..x} f - (x - y) *\<^sub>R f x) {y..x}"
 | |
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 3103 | apply (rule has_integral_diff) | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 3104 | using x y apply (auto intro: integrable_integral [OF integrable_subinterval_real [OF f]]) | 
| 61204 | 3105 | using has_integral_const_real [of "f x" y x] True | 
| 66708 | 3106 | apply simp | 
| 61204 | 3107 | done | 
| 3108 |       have "norm (integral {a..x} f - integral {a..y} f - (x - y) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
 | |
| 3109 | using True | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3110 | apply (simp add: abs_eq_content del: content_real_if measure_lborel_Icc) | 
| 61204 | 3111 | apply (rule has_integral_bound_real[where f="(\<lambda>u. f u - f x)"]) | 
| 66708 | 3112 | using yx True d x y \<open>e>0\<close> assms by (auto simp: Idiff fux_int) | 
| 61204 | 3113 | then show ?thesis | 
| 3114 | by (simp add: algebra_simps norm_minus_commute) | |
| 3115 | qed | |
| 66708 | 3116 |     then have "\<exists>d>0. \<forall>y\<in>{a..b} - S. \<bar>y - x\<bar> < d \<longrightarrow> norm (integral {a..y} f - integral {a..x} f - (y-x) *\<^sub>R f x) \<le> e * \<bar>y - x\<bar>"
 | 
| 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: 
61524diff
changeset | 3117 | using \<open>d>0\<close> by blast | 
| 
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: 
61524diff
changeset | 3118 | } | 
| 61204 | 3119 | then show ?thesis | 
| 3120 | by (simp add: has_vector_derivative_def has_derivative_within_alt bounded_linear_scaleR_left) | |
| 3121 | qed | |
| 3122 | ||
| 66708 | 3123 | |
| 53520 | 3124 | lemma integral_has_vector_derivative: | 
| 3125 | fixes f :: "real \<Rightarrow> 'a::banach" | |
| 66402 | 3126 |   assumes "continuous_on {a..b} f"
 | 
| 3127 |     and "x \<in> {a..b}"
 | |
| 3128 |   shows "((\<lambda>u. integral {a..u} f) has_vector_derivative f(x)) (at x within {a..b})"
 | |
| 66708 | 3129 | using assms integral_has_vector_derivative_continuous_at [OF integrable_continuous_real] | 
| 3130 | by (fastforce simp: continuous_on_eq_continuous_within) | |
| 53520 | 3131 | |
| 3132 | lemma antiderivative_continuous: | |
| 3133 | fixes q b :: real | |
| 66402 | 3134 |   assumes "continuous_on {a..b} f"
 | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 3135 |   obtains g where "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_vector_derivative (f x::_::banach)) (at x within {a..b})"
 | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 3136 | using integral_has_vector_derivative[OF assms] by auto | 
| 53520 | 3137 | |
| 35172 | 3138 | |
| 60420 | 3139 | subsection \<open>Combined fundamental theorem of calculus.\<close> | 
| 35172 | 3140 | |
| 53520 | 3141 | lemma antiderivative_integral_continuous: | 
| 3142 | fixes f :: "real \<Rightarrow> 'a::banach" | |
| 66402 | 3143 |   assumes "continuous_on {a..b} f"
 | 
| 3144 |   obtains g where "\<forall>u\<in>{a..b}. \<forall>v \<in> {a..b}. u \<le> v \<longrightarrow> (f has_integral (g v - g u)) {u..v}"
 | |
| 53520 | 3145 | proof - | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3146 | obtain g | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3147 |     where g: "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_vector_derivative f x) (at x within {a..b})" 
 | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3148 | using antiderivative_continuous[OF assms] by metis | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3149 |   have "(f has_integral g v - g u) {u..v}" if "u \<in> {a..b}" "v \<in> {a..b}" "u \<le> v" for u v
 | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3150 | proof - | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3151 | have "\<And>x. x \<in> cbox u v \<Longrightarrow> (g has_vector_derivative f x) (at x within cbox u v)" | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3152 | by (meson g has_vector_derivative_within_subset interval_subset_is_interval is_interval_closed_interval subsetCE that(1) that(2)) | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3153 | then show ?thesis | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3154 | by (metis box_real(2) that(3) fundamental_theorem_of_calculus) | 
| 53520 | 3155 | qed | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3156 | then show ?thesis | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3157 | using that by blast | 
| 53520 | 3158 | qed | 
| 3159 | ||
| 35172 | 3160 | |
| 60420 | 3161 | subsection \<open>General "twiddling" for interval-to-interval function image.\<close> | 
| 35172 | 3162 | |
| 3163 | lemma has_integral_twiddle: | |
| 53520 | 3164 | assumes "0 < r" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3165 | and hg: "\<And>x. h(g x) = x" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3166 | and gh: "\<And>x. g(h x) = x" | 
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 3167 | and contg: "\<And>x. continuous (at x) g" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3168 | and g: "\<And>u v. \<exists>w z. g ` cbox u v = cbox w z" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3169 | and h: "\<And>u v. \<exists>w z. h ` cbox u v = cbox w z" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3170 | and r: "\<And>u v. content(g ` cbox u v) = r * content (cbox u v)" | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3171 | and intfi: "(f has_integral i) (cbox a b)" | 
| 56188 | 3172 | shows "((\<lambda>x. f(g x)) has_integral (1 / r) *\<^sub>R i) (h ` cbox a b)" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3173 | proof (cases "cbox a b = {}")
 | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3174 | case True | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3175 | then show ?thesis | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3176 | using intfi by auto | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3177 | next | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3178 | case False | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3179 | obtain w z where wz: "h ` cbox a b = cbox w z" | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3180 | using h by blast | 
| 53520 | 3181 | have inj: "inj g" "inj h" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3182 | using hg gh injI by metis+ | 
| 63944 
21eaff8c8fc9
use filter to define Henstock-Kurzweil integration
 hoelzl parents: 
63941diff
changeset | 3183 | from h obtain ha hb where h_eq: "h ` cbox a b = cbox ha hb" by blast | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3184 | have "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of h ` cbox a b \<and> d fine p | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3185 | \<longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e)" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3186 | if "e > 0" for e | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3187 | proof - | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3188 | have "e * r > 0" using that \<open>0 < r\<close> by simp | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3189 | with intfi[unfolded has_integral] | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 3190 | obtain d where "gauge d" | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 3191 | and d: "\<And>p. p tagged_division_of cbox a b \<and> d fine p | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3192 | \<Longrightarrow> norm ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f x) - i) < e * r" | 
| 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3193 | by metis | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 3194 | define d' where "d' x = g -` d (g x)" for x | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3195 | show ?thesis | 
| 53520 | 3196 | proof (rule_tac x=d' in exI, safe) | 
| 3197 | show "gauge d'" | |
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 3198 | using \<open>gauge d\<close> continuous_open_vimage[OF _ contg] by (auto simp: gauge_def d'_def) | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3199 | next | 
| 53520 | 3200 | fix p | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3201 | assume ptag: "p tagged_division_of h ` cbox a b" and finep: "d' fine p" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3202 | note p = tagged_division_ofD[OF ptag] | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3203 | have gab: "g y \<in> cbox a b" if "y \<in> K" "(x, K) \<in> p" for x y K | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3204 | by (metis hg inj(2) inj_image_mem_iff p(3) subsetCE that that) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3205 | have gimp: "(\<lambda>(x,K). (g x, g ` K)) ` p tagged_division_of (cbox a b) \<and> | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3206 | d fine (\<lambda>(x, k). (g x, g ` k)) ` p" | 
| 53520 | 3207 | unfolding tagged_division_of | 
| 3208 | proof safe | |
| 3209 | show "finite ((\<lambda>(x, k). (g x, g ` k)) ` p)" | |
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3210 | using ptag by auto | 
| 53520 | 3211 | show "d fine (\<lambda>(x, k). (g x, g ` k)) ` p" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 3212 | using finep unfolding fine_def d'_def by auto | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3213 | next | 
| 53520 | 3214 | fix x k | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3215 | assume xk: "(x, k) \<in> p" | 
| 53520 | 3216 | show "g x \<in> g ` k" | 
| 3217 | using p(2)[OF xk] by auto | |
| 56188 | 3218 | show "\<exists>u v. g ` k = cbox u v" | 
| 53520 | 3219 | using p(4)[OF xk] using assms(5-6) by auto | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3220 | fix x' K' u | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3221 | assume xk': "(x', K') \<in> p" and u: "u \<in> interior (g ` k)" "u \<in> interior (g ` K')" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3222 |         have "interior k \<inter> interior K' \<noteq> {}"
 | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3223 | proof | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3224 |           assume "interior k \<inter> interior K' = {}"
 | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3225 | moreover have "u \<in> g ` (interior k \<inter> interior K')" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3226 | using interior_image_subset[OF \<open>inj g\<close> contg] u | 
| 63018 
ae2ec7d86ad4
tidying some proofs; getting rid of "nonempty_witness"
 paulson <lp15@cam.ac.uk> parents: 
63007diff
changeset | 3227 | unfolding image_Int[OF inj(1)] by blast | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3228 | ultimately show False by blast | 
| 53520 | 3229 | qed | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3230 | then have same: "(x, k) = (x', K')" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3231 | using ptag xk' xk by blast | 
| 53520 | 3232 | then show "g x = g x'" | 
| 3233 | by auto | |
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3234 | show "g u \<in> g ` K'"if "u \<in> k" for u | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3235 | using that same by auto | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3236 | show "g u \<in> g ` k"if "u \<in> K'" for u | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3237 | using that same by auto | 
| 53520 | 3238 | next | 
| 3239 | fix x | |
| 56188 | 3240 | assume "x \<in> cbox a b" | 
| 53520 | 3241 |         then have "h x \<in>  \<Union>{k. \<exists>x. (x, k) \<in> p}"
 | 
| 3242 | using p(6) by auto | |
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3243 | then obtain X y where "h x \<in> X" "(y, X) \<in> p" by blast | 
| 53520 | 3244 |         then show "x \<in> \<Union>{k. \<exists>x. (x, k) \<in> (\<lambda>(x, k). (g x, g ` k)) ` p}"
 | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3245 | apply clarsimp | 
| 66355 
c828efcb95f3
towards a cleanup of Henstock_Kurzweil_Integration.thy
 paulson <lp15@cam.ac.uk> parents: 
66299diff
changeset | 3246 | by (metis (no_types, lifting) assms(3) image_eqI pair_imageI) | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3247 | qed (use gab in auto) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3248 | have *: "inj_on (\<lambda>(x, k). (g x, g ` k)) p" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3249 | using inj(1) unfolding inj_on_def by fastforce | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3250 | have "(\<Sum>(x, k)\<in>(\<lambda>(x, k). (g x, g ` k)) ` p. content k *\<^sub>R f x) - i = r *\<^sub>R (\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - i" (is "?l = _") | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3251 | using r | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3252 | apply (simp only: algebra_simps add_left_cancel scaleR_right.sum) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3253 | apply (subst sum.reindex_bij_betw[symmetric, where h="\<lambda>(x, k). (g x, g ` k)" and S=p]) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3254 | apply (auto intro!: * sum.cong simp: bij_betw_def dest!: p(4)) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3255 | done | 
| 53520 | 3256 | also have "\<dots> = r *\<^sub>R ((\<Sum>(x, k)\<in>p. content k *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i)" (is "_ = ?r") | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3257 | using \<open>0 < r\<close> by (auto simp: scaleR_diff_right) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3258 | finally have eq: "?l = ?r" . | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3259 | show "norm ((\<Sum>(x,K)\<in>p. content K *\<^sub>R f (g x)) - (1 / r) *\<^sub>R i) < e" | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 3260 | using d[OF gimp] \<open>0 < r\<close> by (auto simp add: eq) | 
| 53520 | 3261 | qed | 
| 3262 | qed | |
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3263 | then show ?thesis | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 3264 | by (auto simp: h_eq has_integral) | 
| 53520 | 3265 | qed | 
| 3266 | ||
| 35172 | 3267 | |
| 60420 | 3268 | subsection \<open>Special case of a basic affine transformation.\<close> | 
| 35172 | 3269 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3270 | lemma AE_lborel_inner_neq: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3271 | assumes k: "k \<in> Basis" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3272 | shows "AE x in lborel. x \<bullet> k \<noteq> c" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3273 | proof - | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3274 | interpret finite_product_sigma_finite "\<lambda>_. lborel" Basis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3275 | proof qed simp | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3276 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3277 |   have "emeasure lborel {x\<in>space lborel. x \<bullet> k = c} = emeasure (\<Pi>\<^sub>M j::'a\<in>Basis. lborel) (\<Pi>\<^sub>E j\<in>Basis. if j = k then {c} else UNIV)"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3278 | using k | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3279 | by (auto simp add: lborel_eq[where 'a='a] emeasure_distr intro!: arg_cong2[where f=emeasure]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3280 | (auto simp: space_PiM PiE_iff extensional_def split: if_split_asm) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3281 |   also have "\<dots> = (\<Prod>j\<in>Basis. emeasure lborel (if j = k then {c} else UNIV))"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3282 | by (intro measure_times) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3283 | also have "\<dots> = 0" | 
| 64272 | 3284 | by (intro prod_zero bexI[OF _ k]) auto | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3285 | finally show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3286 | by (subst AE_iff_measurable[OF _ refl]) auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3287 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3288 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3289 | lemma content_image_stretch_interval: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3290 | fixes m :: "'a::euclidean_space \<Rightarrow> real" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3291 | defines "s f x \<equiv> (\<Sum>k::'a\<in>Basis. (f k * (x\<bullet>k)) *\<^sub>R k)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3292 | shows "content (s m ` cbox a b) = \<bar>\<Prod>k\<in>Basis. m k\<bar> * content (cbox a b)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3293 | proof cases | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3294 | have s[measurable]: "s f \<in> borel \<rightarrow>\<^sub>M borel" for f | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3295 | by (auto simp: s_def[abs_def]) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3296 | assume m: "\<forall>k\<in>Basis. m k \<noteq> 0" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3297 | then have s_comp_s: "s (\<lambda>k. 1 / m k) \<circ> s m = id" "s m \<circ> s (\<lambda>k. 1 / m k) = id" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3298 | by (auto simp: s_def[abs_def] fun_eq_iff euclidean_representation) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3299 | then have "inv (s (\<lambda>k. 1 / m k)) = s m" "bij (s (\<lambda>k. 1 / m k))" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3300 | by (auto intro: inv_unique_comp o_bij) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3301 | then have eq: "s m ` cbox a b = s (\<lambda>k. 1 / m k) -` cbox a b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3302 | using bij_vimage_eq_inv_image[OF \<open>bij (s (\<lambda>k. 1 / m k))\<close>, of "cbox a b"] by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3303 | show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3304 | using m unfolding eq measure_def | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3305 | by (subst lborel_affine_euclidean[where c=m and t=0]) | 
| 64272 | 3306 | (simp_all add: emeasure_density measurable_sets_borel[OF s] abs_prod nn_integral_cmult | 
| 3307 | s_def[symmetric] emeasure_distr vimage_comp s_comp_s enn2real_mult prod_nonneg) | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3308 | next | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3309 | assume "\<not> (\<forall>k\<in>Basis. m k \<noteq> 0)" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3310 | then obtain k where k: "k \<in> Basis" "m k = 0" by auto | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3311 | then have [simp]: "(\<Prod>k\<in>Basis. m k) = 0" | 
| 64272 | 3312 | by (intro prod_zero) auto | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3313 |   have "emeasure lborel {x\<in>space lborel. x \<in> s m ` cbox a b} = 0"
 | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3314 | proof (rule emeasure_eq_0_AE) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3315 | show "AE x in lborel. x \<notin> s m ` cbox a b" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3316 | using AE_lborel_inner_neq[OF \<open>k\<in>Basis\<close>] | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3317 | proof eventually_elim | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3318 | show "x \<bullet> k \<noteq> 0 \<Longrightarrow> x \<notin> s m ` cbox a b " for x | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3319 | using k by (auto simp: s_def[abs_def] cbox_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3320 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3321 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3322 | then show ?thesis | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3323 | by (simp add: measure_def) | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3324 | qed | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 3325 | |
| 53520 | 3326 | lemma interval_image_affinity_interval: | 
| 56188 | 3327 | "\<exists>u v. (\<lambda>x. m *\<^sub>R (x::'a::euclidean_space) + c) ` cbox a b = cbox u v" | 
| 3328 | unfolding image_affinity_cbox | |
| 53520 | 3329 | by auto | 
| 3330 | ||
| 56188 | 3331 | lemma content_image_affinity_cbox: | 
| 3332 | "content((\<lambda>x::'a::euclidean_space. m *\<^sub>R x + c) ` cbox a b) = | |
| 61945 | 3333 |     \<bar>m\<bar> ^ DIM('a) * content (cbox a b)" (is "?l = ?r")
 | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 3334 | proof (cases "cbox a b = {}")
 | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 3335 | case True then show ?thesis by simp | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 3336 | next | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 3337 | case False | 
| 53399 | 3338 | show ?thesis | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3339 | proof (cases "m \<ge> 0") | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3340 | case True | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 3341 |     with \<open>cbox a b \<noteq> {}\<close> have "cbox (m *\<^sub>R a + c) (m *\<^sub>R b + c) \<noteq> {}"
 | 
| 56188 | 3342 | unfolding box_ne_empty | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3343 | apply (intro ballI) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3344 | apply (erule_tac x=i in ballE) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 3345 | apply (auto simp: inner_simps mult_left_mono) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3346 | done | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3347 | moreover from True have *: "\<And>i. (m *\<^sub>R b + c) \<bullet> i - (m *\<^sub>R a + c) \<bullet> i = m *\<^sub>R (b-a) \<bullet> i" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3348 | by (simp add: inner_simps field_simps) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3349 | ultimately show ?thesis | 
| 56188 | 3350 | by (simp add: image_affinity_cbox True content_cbox' | 
| 64272 | 3351 | prod.distrib prod_constant inner_diff_left) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3352 | next | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3353 | case False | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 3354 |     with \<open>cbox a b \<noteq> {}\<close> have "cbox (m *\<^sub>R b + c) (m *\<^sub>R a + c) \<noteq> {}"
 | 
| 56188 | 3355 | unfolding box_ne_empty | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3356 | apply (intro ballI) | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3357 | apply (erule_tac x=i in ballE) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62463diff
changeset | 3358 | apply (auto simp: inner_simps mult_left_mono) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3359 | done | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3360 | moreover from False have *: "\<And>i. (m *\<^sub>R a + c) \<bullet> i - (m *\<^sub>R b + c) \<bullet> i = (-m) *\<^sub>R (b-a) \<bullet> i" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3361 | by (simp add: inner_simps field_simps) | 
| 53374 
a14d2a854c02
tuned proofs -- clarified flow of facts wrt. calculation;
 wenzelm parents: 
53015diff
changeset | 3362 | ultimately show ?thesis using False | 
| 56188 | 3363 | by (simp add: image_affinity_cbox content_cbox' | 
| 64272 | 3364 | prod.distrib[symmetric] prod_constant[symmetric] inner_diff_left) | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3365 | qed | 
| 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3366 | qed | 
| 35172 | 3367 | |
| 53520 | 3368 | lemma has_integral_affinity: | 
| 56188 | 3369 | fixes a :: "'a::euclidean_space" | 
| 3370 | assumes "(f has_integral i) (cbox a b)" | |
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 3371 | and "m \<noteq> 0" | 
| 61945 | 3372 |   shows "((\<lambda>x. f(m *\<^sub>R x + c)) has_integral ((1 / (\<bar>m\<bar> ^ DIM('a))) *\<^sub>R i)) ((\<lambda>x. (1 / m) *\<^sub>R x + -((1 / m) *\<^sub>R c)) ` cbox a b)"
 | 
| 53520 | 3373 | apply (rule has_integral_twiddle) | 
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 3374 | using assms | 
| 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 3375 | apply (safe intro!: interval_image_affinity_interval content_image_affinity_cbox) | 
| 53520 | 3376 | apply (rule zero_less_power) | 
| 61165 | 3377 | unfolding scaleR_right_distrib | 
| 53520 | 3378 | apply auto | 
| 3379 | done | |
| 3380 | ||
| 3381 | lemma integrable_affinity: | |
| 56188 | 3382 | assumes "f integrable_on cbox a b" | 
| 53520 | 3383 | and "m \<noteq> 0" | 
| 56188 | 3384 | shows "(\<lambda>x. f(m *\<^sub>R x + c)) integrable_on ((\<lambda>x. (1 / m) *\<^sub>R x + -((1/m) *\<^sub>R c)) ` cbox a b)" | 
| 53520 | 3385 | using assms | 
| 3386 | unfolding integrable_on_def | |
| 3387 | apply safe | |
| 3388 | apply (drule has_integral_affinity) | |
| 3389 | apply auto | |
| 3390 | done | |
| 3391 | ||
| 60800 
7d04351c795a
New material for Cauchy's integral theorem
 paulson <lp15@cam.ac.uk> parents: 
60762diff
changeset | 3392 | lemmas has_integral_affinity01 = has_integral_affinity [of _ _ 0 "1::real", simplified] | 
| 35172 | 3393 | |
| 60420 | 3394 | subsection \<open>Special case of stretching coordinate axes separately.\<close> | 
| 35172 | 3395 | |
| 53523 | 3396 | lemma has_integral_stretch: | 
| 56188 | 3397 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 3398 | assumes "(f has_integral i) (cbox a b)" | |
| 53523 | 3399 | and "\<forall>k\<in>Basis. m k \<noteq> 0" | 
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 3400 | shows "((\<lambda>x. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) has_integral | 
| 64272 | 3401 | ((1/ \<bar>prod m Basis\<bar>) *\<^sub>R i)) ((\<lambda>x. (\<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k)) ` cbox a b)" | 
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 3402 | apply (rule has_integral_twiddle[where f=f]) | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 3403 | unfolding zero_less_abs_iff content_image_stretch_interval | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 3404 | unfolding image_stretch_interval empty_as_interval euclidean_eq_iff[where 'a='a] | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 3405 | using assms | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 3406 | by auto | 
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 3407 | |
| 37489 
44e42d392c6e
Introduce a type class for euclidean spaces, port most lemmas from real^'n to this type class.
 hoelzl parents: 
36899diff
changeset | 3408 | |
| 53523 | 3409 | lemma integrable_stretch: | 
| 56188 | 3410 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::real_normed_vector" | 
| 3411 | assumes "f integrable_on cbox a b" | |
| 53523 | 3412 | and "\<forall>k\<in>Basis. m k \<noteq> 0" | 
| 3413 | shows "(\<lambda>x::'a. f (\<Sum>k\<in>Basis. (m k * (x\<bullet>k))*\<^sub>R k)) integrable_on | |
| 56188 | 3414 | ((\<lambda>x. \<Sum>k\<in>Basis. (1 / m k * (x\<bullet>k))*\<^sub>R k) ` cbox a b)" | 
| 63928 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 3415 | using assms unfolding integrable_on_def | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 3416 | by (force dest: has_integral_stretch) | 
| 
d81fb5b46a5c
new material about topological concepts, etc
 paulson <lp15@cam.ac.uk> parents: 
63918diff
changeset | 3417 | |
| 53523 | 3418 | |
| 60420 | 3419 | subsection \<open>even more special cases.\<close> | 
| 35172 | 3420 | |
| 53523 | 3421 | lemma uminus_interval_vector[simp]: | 
| 56188 | 3422 | fixes a b :: "'a::euclidean_space" | 
| 3423 | shows "uminus ` cbox a b = cbox (-b) (-a)" | |
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 3424 | apply safe | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 3425 | apply (simp add: mem_box(2)) | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 3426 | apply (rule_tac x="-x" in image_eqI) | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 3427 | apply (auto simp add: mem_box) | 
| 53523 | 3428 | done | 
| 3429 | ||
| 3430 | lemma has_integral_reflect_lemma[intro]: | |
| 56188 | 3431 | assumes "(f has_integral i) (cbox a b)" | 
| 3432 | shows "((\<lambda>x. f(-x)) has_integral i) (cbox (-b) (-a))" | |
| 53523 | 3433 | using has_integral_affinity[OF assms, of "-1" 0] | 
| 3434 | by auto | |
| 3435 | ||
| 56188 | 3436 | lemma has_integral_reflect_lemma_real[intro]: | 
| 66402 | 3437 |   assumes "(f has_integral i) {a..b::real}"
 | 
| 56188 | 3438 |   shows "((\<lambda>x. f(-x)) has_integral i) {-b .. -a}"
 | 
| 3439 | using assms | |
| 3440 | unfolding box_real[symmetric] | |
| 3441 | by (rule has_integral_reflect_lemma) | |
| 3442 | ||
| 53523 | 3443 | lemma has_integral_reflect[simp]: | 
| 56188 | 3444 | "((\<lambda>x. f (-x)) has_integral i) (cbox (-b) (-a)) \<longleftrightarrow> (f has_integral i) (cbox a b)" | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 3445 | by (auto dest: has_integral_reflect_lemma) | 
| 35172 | 3446 | |
| 56188 | 3447 | lemma integrable_reflect[simp]: "(\<lambda>x. f(-x)) integrable_on cbox (-b) (-a) \<longleftrightarrow> f integrable_on cbox a b" | 
| 35172 | 3448 | unfolding integrable_on_def by auto | 
| 3449 | ||
| 66402 | 3450 | lemma integrable_reflect_real[simp]: "(\<lambda>x. f(-x)) integrable_on {-b .. -a} \<longleftrightarrow> f integrable_on {a..b::real}"
 | 
| 56188 | 3451 | unfolding box_real[symmetric] | 
| 3452 | by (rule integrable_reflect) | |
| 3453 | ||
| 3454 | lemma integral_reflect[simp]: "integral (cbox (-b) (-a)) (\<lambda>x. f (-x)) = integral (cbox a b) f" | |
| 35172 | 3455 | unfolding integral_def by auto | 
| 3456 | ||
| 66402 | 3457 | lemma integral_reflect_real[simp]: "integral {-b .. -a} (\<lambda>x. f (-x)) = integral {a..b::real} f"
 | 
| 56188 | 3458 | unfolding box_real[symmetric] | 
| 3459 | by (rule integral_reflect) | |
| 3460 | ||
| 53523 | 3461 | |
| 60420 | 3462 | subsection \<open>Stronger form of FCT; quite a tedious proof.\<close> | 
| 35172 | 3463 | |
| 53523 | 3464 | lemma split_minus[simp]: "(\<lambda>(x, k). f x k) x - (\<lambda>(x, k). g x k) x = (\<lambda>(x, k). f x k - g x k) x" | 
| 3465 | by (simp add: split_def) | |
| 35172 | 3466 | |
| 66382 | 3467 | theorem fundamental_theorem_of_calculus_interior: | 
| 53523 | 3468 | fixes f :: "real \<Rightarrow> 'a::real_normed_vector" | 
| 3469 | assumes "a \<le> b" | |
| 66402 | 3470 |     and contf: "continuous_on {a..b} f"
 | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3471 |     and derf: "\<And>x. x \<in> {a <..< b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
 | 
| 66402 | 3472 |   shows "(f' has_integral (f b - f a)) {a..b}"
 | 
| 66382 | 3473 | proof (cases "a = b") | 
| 3474 | case True | |
| 3475 |   then have *: "cbox a b = {b}" "f b - f a = 0"
 | |
| 3476 | by (auto simp add: order_antisym) | |
| 3477 | with True show ?thesis by auto | |
| 3478 | next | |
| 3479 | case False | |
| 3480 | with \<open>a \<le> b\<close> have ab: "a < b" by arith | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3481 | show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3482 | unfolding has_integral_factor_content_real | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3483 | proof (intro allI impI) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3484 | fix e :: real | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3485 | assume e: "e > 0" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3486 | then have eba8: "(e * (b-a)) / 8 > 0" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3487 | using ab by (auto simp add: field_simps) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3488 | note derf_exp = derf[unfolded has_vector_derivative_def has_derivative_at_alt] | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3489 |     have bounded: "\<And>x. x \<in> {a<..<b} \<Longrightarrow> bounded_linear (\<lambda>u. u *\<^sub>R f' x)"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3490 | using derf_exp by simp | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3491 | have "\<forall>x \<in> box a b. \<exists>d>0. \<forall>y. norm (y-x) < d \<longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e/2 * norm (y-x)" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3492 | (is "\<forall>x \<in> box a b. ?Q x") | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3493 | proof | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3494 | fix x assume x: "x \<in> box a b" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3495 | show "?Q x" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3496 | apply (rule allE [where x="e/2", OF derf_exp [THEN conjunct2, of x]]) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3497 | using x e by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3498 | qed | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3499 | from this [unfolded bgauge_existence_lemma] | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3500 | obtain d where d: "\<And>x. 0 < d x" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3501 | "\<And>x y. \<lbrakk>x \<in> box a b; norm (y-x) < d x\<rbrakk> | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 3502 | \<Longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e/2 * norm (y-x)" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3503 | by metis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3504 | have "bounded (f ` cbox a b)" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3505 | using compact_cbox assms by (auto simp: compact_imp_bounded compact_continuous_image) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3506 | then obtain B | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3507 | where "0 < B" and B: "\<And>x. x \<in> f ` cbox a b \<Longrightarrow> norm x \<le> B" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3508 | unfolding bounded_pos by metis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3509 | obtain da where "0 < da" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3510 |       and da: "\<And>c. \<lbrakk>a \<le> c; {a..c} \<subseteq> {a..b}; {a..c} \<subseteq> ball a da\<rbrakk>
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3511 |                           \<Longrightarrow> norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \<le> (e * (b-a)) / 4"
 | 
| 53523 | 3512 | proof - | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3513 |       have "continuous (at a within {a..b}) f"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3514 | using contf continuous_on_eq_continuous_within by force | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3515 | with eba8 obtain k where "0 < k" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3516 |         and k: "\<And>x. \<lbrakk>x \<in> {a..b}; 0 < norm (x-a); norm (x-a) < k\<rbrakk> \<Longrightarrow> norm (f x - f a) < e * (b-a) / 8"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3517 | unfolding continuous_within Lim_within dist_norm by metis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3518 | obtain l where l: "0 < l" "norm (l *\<^sub>R f' a) \<le> e * (b-a) / 8" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3519 | proof (cases "f' a = 0") | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3520 | case True with ab e that show ?thesis by auto | 
| 53523 | 3521 | next | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3522 | case False | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3523 | then show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3524 | apply (rule_tac l="(e * (b-a)) / 8 / norm (f' a)" in that) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3525 | using ab e apply (auto simp add: field_simps) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3526 | done | 
| 53523 | 3527 | qed | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3528 |       have "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b-a) / 4"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3529 |         if "a \<le> c" "{a..c} \<subseteq> {a..b}" and bmin: "{a..c} \<subseteq> ball a (min k l)" for c
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3530 | proof - | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3531 |         have minkl: "\<bar>a - x\<bar> < min k l" if "x \<in> {a..c}" for x
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3532 | using bmin dist_real_def that by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3533 | then have lel: "\<bar>c - a\<bar> \<le> \<bar>l\<bar>" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3534 | using that by force | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3535 | have "norm ((c - a) *\<^sub>R f' a - (f c - f a)) \<le> norm ((c - a) *\<^sub>R f' a) + norm (f c - f a)" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3536 | by (rule norm_triangle_ineq4) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3537 | also have "\<dots> \<le> e * (b-a) / 8 + e * (b-a) / 8" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3538 | proof (rule add_mono) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3539 | have "norm ((c - a) *\<^sub>R f' a) \<le> norm (l *\<^sub>R f' a)" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3540 | by (auto intro: mult_right_mono [OF lel]) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3541 | also have "... \<le> e * (b-a) / 8" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3542 | by (rule l) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3543 | finally show "norm ((c - a) *\<^sub>R f' a) \<le> e * (b-a) / 8" . | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3544 | next | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3545 | have "norm (f c - f a) < e * (b-a) / 8" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3546 | proof (cases "a = c") | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3547 | case True then show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3548 | using eba8 by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3549 | next | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3550 | case False show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3551 | by (rule k) (use minkl \<open>a \<le> c\<close> that False in auto) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3552 | qed | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3553 | then show "norm (f c - f a) \<le> e * (b-a) / 8" by simp | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3554 | qed | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3555 |         finally show "norm (content {a..c} *\<^sub>R f' a - (f c - f a)) \<le> e * (b-a) / 4"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3556 | unfolding content_real[OF \<open>a \<le> c\<close>] by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3557 | qed | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3558 | then show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3559 | by (rule_tac da="min k l" in that) (auto simp: l \<open>0 < k\<close>) | 
| 53523 | 3560 | qed | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3561 | obtain db where "0 < db" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3562 |       and db: "\<And>c. \<lbrakk>c \<le> b; {c..b} \<subseteq> {a..b}; {c..b} \<subseteq> ball b db\<rbrakk>
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3563 |                           \<Longrightarrow> norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \<le> (e * (b-a)) / 4"
 | 
| 53523 | 3564 | proof - | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3565 |       have "continuous (at b within {a..b}) f"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3566 | using contf continuous_on_eq_continuous_within by force | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3567 | with eba8 obtain k | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3568 | where "0 < k" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3569 |           and k: "\<And>x. \<lbrakk>x \<in> {a..b}; 0 < norm(x-b); norm(x-b) < k\<rbrakk>
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3570 | \<Longrightarrow> norm (f b - f x) < e * (b-a) / 8" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3571 | unfolding continuous_within Lim_within dist_norm norm_minus_commute by metis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3572 | obtain l where l: "0 < l" "norm (l *\<^sub>R f' b) \<le> (e * (b-a)) / 8" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3573 | proof (cases "f' b = 0") | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3574 | case True thus ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3575 | using ab e that by auto | 
| 53523 | 3576 | next | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3577 | case False then show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3578 | apply (rule_tac l="(e * (b-a)) / 8 / norm (f' b)" in that) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3579 | using ab e by (auto simp add: field_simps) | 
| 53523 | 3580 | qed | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3581 |       have "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b-a) / 4" 
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3582 |         if "c \<le> b" "{c..b} \<subseteq> {a..b}" and bmin: "{c..b} \<subseteq> ball b (min k l)" for c
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3583 | proof - | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3584 |         have minkl: "\<bar>b - x\<bar> < min k l" if "x \<in> {c..b}" for x
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3585 | using bmin dist_real_def that by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3586 | then have lel: "\<bar>b - c\<bar> \<le> \<bar>l\<bar>" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3587 | using that by force | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3588 | have "norm ((b - c) *\<^sub>R f' b - (f b - f c)) \<le> norm ((b - c) *\<^sub>R f' b) + norm (f b - f c)" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3589 | by (rule norm_triangle_ineq4) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3590 | also have "\<dots> \<le> e * (b-a) / 8 + e * (b-a) / 8" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3591 | proof (rule add_mono) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3592 | have "norm ((b - c) *\<^sub>R f' b) \<le> norm (l *\<^sub>R f' b)" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3593 | by (auto intro: mult_right_mono [OF lel]) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3594 | also have "... \<le> e * (b-a) / 8" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3595 | by (rule l) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3596 | finally show "norm ((b - c) *\<^sub>R f' b) \<le> e * (b-a) / 8" . | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3597 | next | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3598 | have "norm (f b - f c) < e * (b-a) / 8" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3599 | proof (cases "b = c") | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3600 | case True with eba8 show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3601 | by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3602 | next | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3603 | case False show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3604 | by (rule k) (use minkl \<open>c \<le> b\<close> that False in auto) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3605 | qed | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3606 | then show "norm (f b - f c) \<le> e * (b-a) / 8" by simp | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3607 | qed | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3608 |         finally show "norm (content {c..b} *\<^sub>R f' b - (f b - f c)) \<le> e * (b-a) / 4"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3609 | unfolding content_real[OF \<open>c \<le> b\<close>] by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3610 | qed | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3611 | then show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3612 | by (rule_tac db="min k l" in that) (auto simp: l \<open>0 < k\<close>) | 
| 53523 | 3613 | qed | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3614 | let ?d = "(\<lambda>x. ball x (if x=a then da else if x=b then db else d x))" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3615 |     show "\<exists>d. gauge d \<and> (\<forall>p. p tagged_division_of {a..b} \<and> d fine p \<longrightarrow>
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3616 |               norm ((\<Sum>(x,K)\<in>p. content K *\<^sub>R f' x) - (f b - f a)) \<le> e * content {a..b})"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3617 | proof (rule exI, safe) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3618 | show "gauge ?d" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3619 | using ab \<open>db > 0\<close> \<open>da > 0\<close> d(1) by (auto intro: gauge_ball_dependent) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3620 | next | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3621 | fix p | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3622 |       assume ptag: "p tagged_division_of {a..b}" and fine: "?d fine p"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3623 |       let ?A = "{t. fst t \<in> {a, b}}"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3624 | note p = tagged_division_ofD[OF ptag] | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3625 |       have pA: "p = (p \<inter> ?A) \<union> (p - ?A)" "finite (p \<inter> ?A)" "finite (p - ?A)" "(p \<inter> ?A) \<inter> (p - ?A) = {}"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3626 | using ptag fine by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3627 | have le_xz: "\<And>w x y z::real. y \<le> z/2 \<Longrightarrow> w - x \<le> z/2 \<Longrightarrow> w + y \<le> x + z" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3628 | by arith | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3629 | have non: False if xk: "(x,K) \<in> p" and "x \<noteq> a" "x \<noteq> b" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3630 | and less: "e * (Sup K - Inf K)/2 < norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3631 | for x K | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3632 | proof - | 
| 66400 | 3633 | obtain u v where k: "K = cbox u v" | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3634 | using p(4) xk by blast | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3635 |         then have "u \<le> v" and uv: "{u, v} \<subseteq> cbox u v"
 | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3636 | using p(2)[OF xk] by auto | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3637 | then have result: "e * (v - u)/2 < norm ((v - u) *\<^sub>R f' x - (f v - f u))" | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3638 | using less[unfolded k box_real interval_bounds_real content_real] by auto | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3639 | then have "x \<in> box a b" | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3640 | using p(2) p(3) \<open>x \<noteq> a\<close> \<open>x \<noteq> b\<close> xk by fastforce | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3641 | with d have *: "\<And>y. norm (y-x) < d x | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 3642 | \<Longrightarrow> norm (f y - f x - (y-x) *\<^sub>R f' x) \<le> e/2 * norm (y-x)" | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3643 | by metis | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3644 | have xd: "norm (u - x) < d x" "norm (v - x) < d x" | 
| 66400 | 3645 | using fineD[OF fine xk] \<open>x \<noteq> a\<close> \<open>x \<noteq> b\<close> uv | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3646 | by (auto simp add: k subset_eq dist_commute dist_real_def) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3647 | have "norm ((v - u) *\<^sub>R f' x - (f v - f u)) = | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3648 | norm ((f u - f x - (u - x) *\<^sub>R f' x) - (f v - f x - (v - x) *\<^sub>R f' x))" | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3649 | by (rule arg_cong[where f=norm]) (auto simp: scaleR_left.diff) | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 3650 | also have "\<dots> \<le> e/2 * norm (u - x) + e/2 * norm (v - x)" | 
| 66420 | 3651 | by (metis norm_triangle_le_diff add_mono * xd) | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 3652 | also have "\<dots> \<le> e/2 * norm (v - u)" | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3653 | using p(2)[OF xk] by (auto simp add: field_simps k) | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3654 | also have "\<dots> < norm ((v - u) *\<^sub>R f' x - (f v - f u))" | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3655 | using result by (simp add: \<open>u \<le> v\<close>) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3656 | finally have "e * (v - u)/2 < e * (v - u)/2" | 
| 66384 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3657 | using uv by auto | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3658 | then show False by auto | 
| 
cc66710c9d48
more cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66383diff
changeset | 3659 | qed | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3660 | have "norm (\<Sum>(x, K)\<in>p - ?A. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) | 
| 66400 | 3661 | \<le> (\<Sum>(x, K)\<in>p - ?A. norm (content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3662 | by (auto intro: sum_norm_le) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3663 | also have "... \<le> (\<Sum>n\<in>p - ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k)/2)" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3664 | using non by (fastforce intro: sum_mono) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3665 | finally have I: "norm (\<Sum>(x, k)\<in>p - ?A. | 
| 66383 
5eb0faf4477a
partly unravelled fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66382diff
changeset | 3666 | content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3667 | \<le> (\<Sum>n\<in>p - ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))/2" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3668 | by (simp add: sum_divide_distrib) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3669 | have II: "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) - | 
| 66383 
5eb0faf4477a
partly unravelled fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66382diff
changeset | 3670 | (\<Sum>n\<in>p \<inter> ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k)) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3671 | \<le> (\<Sum>n\<in>p - ?A. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))/2" | 
| 53523 | 3672 | proof - | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3673 | have ge0: "0 \<le> e * (Sup k - Inf k)" if xkp: "(x, k) \<in> p \<inter> ?A" for x k | 
| 66400 | 3674 | proof - | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3675 | obtain u v where uv: "k = cbox u v" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3676 | by (meson Int_iff xkp p(4)) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3677 |           with p(2) that uv have "cbox u v \<noteq> {}"
 | 
| 66400 | 3678 | by blast | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3679 | then show "0 \<le> e * ((Sup k) - (Inf k))" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3680 | unfolding uv using e by (auto simp add: field_simps) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3681 | qed | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3682 |         let ?B = "\<lambda>x. {t \<in> p. fst t = x \<and> content (snd t) \<noteq> 0}"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3683 |         let ?C = "{t \<in> p. fst t \<in> {a, b} \<and> content (snd t) \<noteq> 0}"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3684 |         have "norm (\<Sum>(x, k)\<in>p \<inter> {t. fst t \<in> {a, b}}. content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \<le> e * (b-a)/2"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3685 | proof - | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3686 | have *: "\<And>S f e. sum f S = sum f (p \<inter> ?C) \<Longrightarrow> norm (sum f (p \<inter> ?C)) \<le> e \<Longrightarrow> norm (sum f S) \<le> e" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3687 | by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3688 | have 1: "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3689 |             if "(x,K) \<in> p \<inter> {t. fst t \<in> {a, b}} - p \<inter> ?C" for x K
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3690 | proof - | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3691 | have xk: "(x,K) \<in> p" and k0: "content K = 0" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3692 | using that by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3693 | then obtain u v where uv: "K = cbox u v" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3694 | using p(4) by blast | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3695 | then have "u = v" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3696 | using xk k0 p(2) by force | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3697 | then show "content K *\<^sub>R (f' x) - (f ((Sup K)) - f ((Inf K))) = 0" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3698 | using xk unfolding uv by auto | 
| 61165 | 3699 | qed | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3700 | have 2: "norm(\<Sum>(x,K)\<in>p \<inter> ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (b-a)/2" | 
| 53523 | 3701 | proof - | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3702 | have norm_le: "norm (sum f S) \<le> e" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3703 | if "\<And>x y. \<lbrakk>x \<in> S; y \<in> S\<rbrakk> \<Longrightarrow> x = y" "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> e" "e > 0" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3704 | for S f and e :: real | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3705 |             proof (cases "S = {}")
 | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3706 | case True | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3707 | with that show ?thesis by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3708 | next | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3709 | case False then obtain x where "x \<in> S" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3710 | by auto | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3711 |               then have "S = {x}"
 | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3712 | using that(1) by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3713 | then show ?thesis | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3714 | using \<open>x \<in> S\<close> that(2) by auto | 
| 35172 | 3715 | qed | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3716 | have *: "p \<inter> ?C = ?B a \<union> ?B b" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3717 | by blast | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3718 | then have "norm (\<Sum>(x,K)\<in>p \<inter> ?C. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) = | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3719 | norm (\<Sum>(x,K) \<in> ?B a \<union> ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3720 | by simp | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3721 | also have "... = norm ((\<Sum>(x,K) \<in> ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) + | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3722 | (\<Sum>(x,K) \<in> ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))))" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3723 | apply (subst sum.union_disjoint) | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3724 | using p(1) ab e by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3725 | also have "... \<le> e * (b - a) / 4 + e * (b - a) / 4" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3726 | proof (rule norm_triangle_le [OF add_mono]) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3727 | have pa: "\<exists>v. k = cbox a v \<and> a \<le> v" if "(a, k) \<in> p" for k | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3728 | using p(2) p(3) p(4) that by fastforce | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3729 | show "norm (\<Sum>(x,K) \<in> ?B a. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (b - a) / 4" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3730 | proof (intro norm_le; clarsimp) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3731 | fix K K' | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3732 | assume K: "(a, K) \<in> p" "(a, K') \<in> p" and ne0: "content K \<noteq> 0" "content K' \<noteq> 0" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3733 | with pa obtain v v' where v: "K = cbox a v" "a \<le> v" and v': "K' = cbox a v'" "a \<le> v'" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3734 | by blast | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3735 | let ?v = "min v v'" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3736 | have "box a ?v \<subseteq> K \<inter> K'" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3737 | unfolding v v' by (auto simp add: mem_box) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3738 | then have "interior (box a (min v v')) \<subseteq> interior K \<inter> interior K'" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3739 | using interior_Int interior_mono by blast | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3740 | moreover have "(a + ?v)/2 \<in> box a ?v" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3741 | using ne0 unfolding v v' content_eq_0 not_le | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3742 | by (auto simp add: mem_box) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3743 | ultimately have "(a + ?v)/2 \<in> interior K \<inter> interior K'" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3744 | unfolding interior_open[OF open_box] by auto | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3745 | then show "K = K'" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3746 | using p(5)[OF K] by auto | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3747 | next | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3748 | fix K | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3749 | assume K: "(a, K) \<in> p" and ne0: "content K \<noteq> 0" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3750 | show "norm (content c *\<^sub>R f' a - (f (Sup c) - f (Inf c))) * 4 \<le> e * (b-a)" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3751 | if "(a, c) \<in> p" and ne0: "content c \<noteq> 0" for c | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3752 | proof - | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3753 | obtain v where v: "c = cbox a v" and "a \<le> v" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3754 | using pa[OF \<open>(a, c) \<in> p\<close>] by metis | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3755 |                   then have "a \<in> {a..v}" "v \<le> b"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3756 | using p(3)[OF \<open>(a, c) \<in> p\<close>] by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3757 |                   moreover have "{a..v} \<subseteq> ball a da"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3758 | using fineD[OF \<open>?d fine p\<close> \<open>(a, c) \<in> p\<close>] by (simp add: v split: if_split_asm) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3759 | ultimately show ?thesis | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3760 | unfolding v interval_bounds_real[OF \<open>a \<le> v\<close>] box_real | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3761 | using da \<open>a \<le> v\<close> by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3762 | qed | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3763 | qed (use ab e in auto) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3764 | next | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3765 | have pb: "\<exists>v. k = cbox v b \<and> b \<ge> v" if "(b, k) \<in> p" for k | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3766 | using p(2) p(3) p(4) that by fastforce | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3767 | show "norm (\<Sum>(x,K) \<in> ?B b. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) \<le> e * (b - a) / 4" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3768 | proof (intro norm_le; clarsimp) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3769 | fix K K' | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3770 | assume K: "(b, K) \<in> p" "(b, K') \<in> p" and ne0: "content K \<noteq> 0" "content K' \<noteq> 0" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3771 | with pb obtain v v' where v: "K = cbox v b" "v \<le> b" and v': "K' = cbox v' b" "v' \<le> b" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3772 | by blast | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3773 | let ?v = "max v v'" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3774 | have "box ?v b \<subseteq> K \<inter> K'" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3775 | unfolding v v' by (auto simp: mem_box) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3776 | then have "interior (box (max v v') b) \<subseteq> interior K \<inter> interior K'" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3777 | using interior_Int interior_mono by blast | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3778 | moreover have " ((b + ?v)/2) \<in> box ?v b" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3779 | using ne0 unfolding v v' content_eq_0 not_le by (auto simp: mem_box) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3780 | ultimately have "((b + ?v)/2) \<in> interior K \<inter> interior K'" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3781 | unfolding interior_open[OF open_box] by auto | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3782 | then show "K = K'" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3783 | using p(5)[OF K] by auto | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3784 | next | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3785 | fix K | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3786 | assume K: "(b, K) \<in> p" and ne0: "content K \<noteq> 0" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3787 | show "norm (content c *\<^sub>R f' b - (f (Sup c) - f (Inf c))) * 4 \<le> e * (b-a)" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3788 | if "(b, c) \<in> p" and ne0: "content c \<noteq> 0" for c | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3789 | proof - | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3790 | obtain v where v: "c = cbox v b" and "v \<le> b" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3791 | using \<open>(b, c) \<in> p\<close> pb by blast | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3792 |                 then have "v \<ge> a""b \<in> {v.. b}"  
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3793 | using p(3)[OF \<open>(b, c) \<in> p\<close>] by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3794 |                 moreover have "{v..b} \<subseteq> ball b db"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3795 | using fineD[OF \<open>?d fine p\<close> \<open>(b, c) \<in> p\<close>] box_real(2) v False by force | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3796 | ultimately show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3797 | using db v by auto | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3798 | qed | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3799 | qed (use ab e in auto) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3800 | qed | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3801 | also have "... = e * (b-a)/2" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3802 | by simp | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3803 | finally show "norm (\<Sum>(x,k)\<in>p \<inter> ?C. | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3804 | content k *\<^sub>R f' x - (f (Sup k) - f (Inf k))) \<le> e * (b-a)/2" . | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3805 | qed | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3806 | show "norm (\<Sum>(x, k)\<in>p \<inter> ?A. content k *\<^sub>R f' x - (f ((Sup k)) - f ((Inf k)))) \<le> e * (b-a)/2" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3807 | apply (rule * [OF sum.mono_neutral_right[OF pA(2)]]) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3808 | using 1 2 by (auto simp: split_paired_all) | 
| 66400 | 3809 | qed | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3810 | also have "... = (\<Sum>n\<in>p. e * (case n of (x, k) \<Rightarrow> Sup k - Inf k))/2" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3811 | unfolding sum_distrib_left[symmetric] | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3812 | apply (subst additive_tagged_division_1[OF \<open>a \<le> b\<close> ptag]) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3813 | by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3814 |         finally have norm_le: "norm (\<Sum>(x,K)\<in>p \<inter> {t. fst t \<in> {a, b}}. content K *\<^sub>R f' x - (f (Sup K) - f (Inf K)))
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3815 | \<le> (\<Sum>n\<in>p. e * (case n of (x, K) \<Rightarrow> Sup K - Inf K))/2" . | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3816 | have le2: "\<And>x s1 s2::real. 0 \<le> s1 \<Longrightarrow> x \<le> (s1 + s2)/2 \<Longrightarrow> x - s1 \<le> s2/2" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3817 | by auto | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3818 | show ?thesis | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3819 | apply (rule le2 [OF sum_nonneg]) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3820 | using ge0 apply force | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3821 | unfolding sum.union_disjoint[OF pA(2-), symmetric] pA(1)[symmetric] | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3822 | by (metis norm_le) | 
| 53523 | 3823 | qed | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3824 | note * = additive_tagged_division_1[OF assms(1) ptag, symmetric] | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3825 | have "norm (\<Sum>(x,K)\<in>p \<inter> ?A \<union> (p - ?A). content K *\<^sub>R f' x - (f (Sup K) - f (Inf K))) | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3826 | \<le> e * (\<Sum>(x,K)\<in>p \<inter> ?A \<union> (p - ?A). Sup K - Inf K)" | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3827 | unfolding sum_distrib_left | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3828 | unfolding sum.union_disjoint[OF pA(2-)] | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3829 | using le_xz norm_triangle_le I II by blast | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3830 | then | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3831 |       show "norm ((\<Sum>(x,K)\<in>p. content K *\<^sub>R f' x) - (f b - f a)) \<le> e * content {a..b}"
 | 
| 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3832 | by (simp only: content_real[OF \<open>a \<le> b\<close>] *[of "\<lambda>x. x"] *[of f] sum_subtractf[symmetric] split_minus pA(1) [symmetric]) | 
| 53523 | 3833 | qed | 
| 3834 | qed | |
| 3835 | qed | |
| 3836 | ||
| 35172 | 3837 | |
| 60420 | 3838 | subsection \<open>Stronger form with finite number of exceptional points.\<close> | 
| 35172 | 3839 | |
| 53524 | 3840 | lemma fundamental_theorem_of_calculus_interior_strong: | 
| 3841 | fixes f :: "real \<Rightarrow> 'a::banach" | |
| 3842 | assumes "finite s" | |
| 3843 | and "a \<le> b" | |
| 66402 | 3844 |     and "continuous_on {a..b} f"
 | 
| 56188 | 3845 |     and "\<forall>x\<in>{a <..< b} - s. (f has_vector_derivative f'(x)) (at x)"
 | 
| 66402 | 3846 |   shows "(f' has_integral (f b - f a)) {a..b}"
 | 
| 53524 | 3847 | using assms | 
| 3848 | proof (induct "card s" arbitrary: s a b) | |
| 3849 | case 0 | |
| 66402 | 3850 | then show ?case | 
| 3851 | by (force simp add: intro: fundamental_theorem_of_calculus_interior) | |
| 53524 | 3852 | next | 
| 3853 | case (Suc n) | |
| 66402 | 3854 | then obtain c s' where cs: "s = insert c s'" and n: "n = card s'" | 
| 3855 | by (metis card_eq_SucD) | |
| 3856 | then have "finite s'" | |
| 3857 | using \<open>finite s\<close> by force | |
| 53524 | 3858 | show ?case | 
| 54775 
2d3df8633dad
prefer box over greaterThanLessThan on euclidean_space
 immler parents: 
54411diff
changeset | 3859 | proof (cases "c \<in> box a b") | 
| 53524 | 3860 | case False | 
| 66402 | 3861 | with \<open>finite s'\<close> show ?thesis | 
| 3862 | using cs n Suc | |
| 3863 | by (metis Diff_iff box_real(1) insert_iff) | |
| 53524 | 3864 | next | 
| 66402 | 3865 |     let ?P = "\<lambda>i j. \<forall>x\<in>{i <..< j} - s'. (f has_vector_derivative f' x) (at x)"
 | 
| 53524 | 3866 | case True | 
| 3867 | then have "a \<le> c" "c \<le> b" | |
| 56188 | 3868 | by (auto simp: mem_box) | 
| 66402 | 3869 | moreover have "?P a c" "?P c b" | 
| 3870 | using Suc.prems(4) True \<open>a \<le> c\<close> cs(1) by auto | |
| 3871 |     moreover have "continuous_on {a..c} f" "continuous_on {c..b} f"
 | |
| 3872 |       using \<open>continuous_on {a..b} f\<close> \<open>a \<le> c\<close> \<open>c \<le> b\<close> continuous_on_subset by fastforce+
 | |
| 3873 |     ultimately have "(f' has_integral f c - f a + (f b - f c)) {a..b}"
 | |
| 3874 | using Suc.hyps(1) \<open>finite s'\<close> \<open>n = card s'\<close> by (blast intro: has_integral_combine) | |
| 3875 | then show ?thesis | |
| 3876 | by auto | |
| 53524 | 3877 | qed | 
| 3878 | qed | |
| 3879 | ||
| 66402 | 3880 | corollary fundamental_theorem_of_calculus_strong: | 
| 53524 | 3881 | fixes f :: "real \<Rightarrow> 'a::banach" | 
| 3882 | assumes "finite s" | |
| 3883 | and "a \<le> b" | |
| 66402 | 3884 |     and "continuous_on {a..b} f"
 | 
| 3885 |     and vec: "\<forall>x\<in>{a..b} - s. (f has_vector_derivative f'(x)) (at x)"
 | |
| 3886 |   shows "(f' has_integral (f b - f a)) {a..b}"
 | |
| 53524 | 3887 | apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1-3), of f']) | 
| 66402 | 3888 | using vec apply (auto simp: mem_box) | 
| 53524 | 3889 | done | 
| 3890 | ||
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3891 | proposition indefinite_integral_continuous_left: | 
| 53634 | 3892 | fixes f:: "real \<Rightarrow> 'a::banach" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3893 |   assumes intf: "f integrable_on {a..b}" and "a < c" "c \<le> b" "e > 0"
 | 
| 53634 | 3894 | obtains d where "d > 0" | 
| 66402 | 3895 |     and "\<forall>t. c - d < t \<and> t \<le> c \<longrightarrow> norm (integral {a..c} f - integral {a..t} f) < e"
 | 
| 53634 | 3896 | proof - | 
| 66402 | 3897 | obtain w where "w > 0" and w: "\<And>t. \<lbrakk>c - w < t; t < c\<rbrakk> \<Longrightarrow> norm (f c) * norm(c - t) < e/3" | 
| 53634 | 3898 | proof (cases "f c = 0") | 
| 3899 | case False | |
| 66402 | 3900 | hence e3: "0 < e/3 / norm (f c)" using \<open>e>0\<close> by simp | 
| 3901 | moreover have "norm (f c) * norm (c - t) < e/3" | |
| 3902 | if "t < c" and "c - e/3 / norm (f c) < t" for t | |
| 53634 | 3903 | proof - | 
| 66402 | 3904 | have "norm (c - t) < e/3 / norm (f c)" | 
| 3905 | using that by auto | |
| 3906 | then show "norm (f c) * norm (c - t) < e/3" | |
| 3907 | by (metis e3 mult.commute norm_not_less_zero pos_less_divide_eq zero_less_divide_iff) | |
| 53634 | 3908 | qed | 
| 66402 | 3909 | ultimately show ?thesis | 
| 3910 | using that by auto | |
| 53634 | 3911 | next | 
| 66402 | 3912 | case True then show ?thesis | 
| 3913 | using \<open>e > 0\<close> that by auto | |
| 53634 | 3914 | qed | 
| 66402 | 3915 | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3916 | let ?SUM = "\<lambda>p. (\<Sum>(x,K) \<in> p. content K *\<^sub>R f x)" | 
| 66402 | 3917 | have e3: "e/3 > 0" | 
| 3918 | using \<open>e > 0\<close> by auto | |
| 3919 |   have "f integrable_on {a..c}"
 | |
| 3920 | apply (rule integrable_subinterval_real[OF intf]) | |
| 3921 | using \<open>a < c\<close> \<open>c \<le> b\<close> by auto | |
| 3922 | then obtain d1 where "gauge d1" and d1: | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3923 |     "\<And>p. \<lbrakk>p tagged_division_of {a..c}; d1 fine p\<rbrakk> \<Longrightarrow> norm (?SUM p - integral {a..c} f) < e/3"
 | 
| 66402 | 3924 | using integrable_integral has_integral_real e3 by metis | 
| 63040 | 3925 | define d where [abs_def]: "d x = ball x w \<inter> d1 x" for x | 
| 53634 | 3926 | have "gauge d" | 
| 66402 | 3927 | unfolding d_def using \<open>w > 0\<close> \<open>gauge d1\<close> by auto | 
| 3928 | then obtain k where "0 < k" and k: "ball c k \<subseteq> d c" | |
| 3929 | by (meson gauge_def open_contains_ball) | |
| 53634 | 3930 | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3931 | let ?d = "min k (c - a)/2" | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3932 | show thesis | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 3933 | proof (intro that[of ?d] allI impI, safe) | 
| 53634 | 3934 | show "?d > 0" | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3935 | using \<open>0 < k\<close> \<open>a < c\<close> by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3936 | next | 
| 53634 | 3937 | fix t | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3938 | assume t: "c - ?d < t" "t \<le> c" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3939 |     show "norm (integral ({a..c}) f - integral ({a..t}) f) < e"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3940 | proof (cases "t < c") | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3941 | case False with \<open>t \<le> c\<close> show ?thesis | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3942 | by (simp add: \<open>e > 0\<close>) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3943 | next | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3944 | case True | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3945 |       have "f integrable_on {a..t}"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3946 | apply (rule integrable_subinterval_real[OF intf]) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3947 | using \<open>t < c\<close> \<open>c \<le> b\<close> by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3948 | then obtain d2 where d2: "gauge d2" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3949 |         "\<And>p. p tagged_division_of {a..t} \<and> d2 fine p \<Longrightarrow> norm (?SUM p - integral {a..t} f) < e/3"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3950 | using integrable_integral has_integral_real e3 by metis | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3951 | define d3 where "d3 x = (if x \<le> t then d1 x \<inter> d2 x else d1 x)" for x | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3952 | have "gauge d3" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3953 | using \<open>gauge d1\<close> \<open>gauge d2\<close> unfolding d3_def gauge_def by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3954 |       then obtain p where ptag: "p tagged_division_of {a..t}" and pfine: "d3 fine p"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3955 | by (metis box_real(2) fine_division_exists) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3956 | note p' = tagged_division_ofD[OF ptag] | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3957 | have pt: "(x,K)\<in>p \<Longrightarrow> x \<le> t" for x K | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3958 | by (meson atLeastAtMost_iff p'(2) p'(3) subsetCE) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3959 | with pfine have "d2 fine p" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3960 | unfolding fine_def d3_def by fastforce | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3961 |       then have d2_fin: "norm (?SUM p - integral {a..t} f) < e/3"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3962 | using d2(2) ptag by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3963 |       have eqs: "{a..c} \<inter> {x. x \<le> t} = {a..t}" "{a..c} \<inter> {x. x \<ge> t} = {t..c}"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3964 | using t by (auto simp add: field_simps) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3965 |       have "p \<union> {(c, {t..c})} tagged_division_of {a..c}"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3966 | apply (rule tagged_division_Un_interval_real[of _ _ _ 1 "t"]) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3967 | using \<open>t \<le> c\<close> by (auto simp: eqs ptag tagged_division_of_self_real) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3968 | moreover | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3969 |       have "d1 fine p \<union> {(c, {t..c})}"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3970 | unfolding fine_def | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3971 | proof safe | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3972 | fix x K y | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3973 | assume "(x,K) \<in> p" and "y \<in> K" then show "y \<in> d1 x" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3974 | by (metis Int_iff d3_def subsetD fineD pfine) | 
| 66402 | 3975 | next | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3976 |         fix x assume "x \<in> {t..c}"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3977 | then have "dist c x < k" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3978 | using t(1) by (auto simp add: field_simps dist_real_def) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3979 | with k show "x \<in> d1 c" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3980 | unfolding d_def by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3981 | qed | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3982 |       ultimately have d1_fin: "norm (?SUM(p \<union> {(c, {t..c})}) - integral {a..c} f) < e/3"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3983 | using d1 by metis | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3984 |       have SUMEQ: "?SUM(p \<union> {(c, {t..c})}) = (c - t) *\<^sub>R f c + ?SUM p"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3985 | proof - | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3986 |         have "?SUM(p \<union> {(c, {t..c})}) = (content{t..c} *\<^sub>R f c) + ?SUM p"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3987 | proof (subst sum.union_disjoint) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3988 |           show "p \<inter> {(c, {t..c})} = {}"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3989 | using \<open>t < c\<close> pt by force | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3990 | qed (use p'(1) in auto) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3991 | also have "... = (c - t) *\<^sub>R f c + ?SUM p" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3992 | using \<open>t \<le> c\<close> by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3993 | finally show ?thesis . | 
| 66402 | 3994 | qed | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3995 | have "c - k < t" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3996 | using \<open>k>0\<close> t(1) by (auto simp add: field_simps) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3997 | moreover have "k \<le> w" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3998 | proof (rule ccontr) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 3999 | assume "\<not> k \<le> w" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4000 | then have "c + (k + w) / 2 \<notin> d c" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4001 | by (auto simp add: field_simps not_le not_less dist_real_def d_def) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4002 | then have "c + (k + w) / 2 \<notin> ball c k" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4003 | using k by blast | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4004 | then show False | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4005 | using \<open>0 < w\<close> \<open>\<not> k \<le> w\<close> dist_real_def by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4006 | qed | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4007 | ultimately have cwt: "c - w < t" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4008 | by (auto simp add: field_simps) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4009 |       have eq: "integral {a..c} f - integral {a..t} f = -(((c - t) *\<^sub>R f c + ?SUM p) -
 | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 4010 |              integral {a..c} f) + (?SUM p - integral {a..t} f) + (c - t) *\<^sub>R f c"
 | 
| 66533 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4011 | by auto | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4012 |       have "norm (integral {a..c} f - integral {a..t} f) < e/3 + e/3 + e/3"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4013 | unfolding eq | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4014 | proof (intro norm_triangle_lt add_strict_mono) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4015 |         show "norm (- ((c - t) *\<^sub>R f c + ?SUM p - integral {a..c} f)) < e/3"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4016 | by (metis SUMEQ d1_fin norm_minus_cancel) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4017 |         show "norm (?SUM p - integral {a..t} f) < e/3"
 | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4018 | using d2_fin by blast | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4019 | show "norm ((c - t) *\<^sub>R f c) < e/3" | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4020 | using w cwt \<open>t < c\<close> by (auto simp add: field_simps) | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4021 | qed | 
| 
c485474cd91d
Giant cleanup of fundamental_theorem_of_calculus_interior
 paulson <lp15@cam.ac.uk> parents: 
66532diff
changeset | 4022 | then show ?thesis by simp | 
| 53634 | 4023 | qed | 
| 4024 | qed | |
| 4025 | qed | |
| 4026 | ||
| 4027 | lemma indefinite_integral_continuous_right: | |
| 4028 | fixes f :: "real \<Rightarrow> 'a::banach" | |
| 66402 | 4029 |   assumes "f integrable_on {a..b}"
 | 
| 53634 | 4030 | and "a \<le> c" | 
| 4031 | and "c < b" | |
| 4032 | and "e > 0" | |
| 4033 | obtains d where "0 < d" | |
| 66402 | 4034 |     and "\<forall>t. c \<le> t \<and> t < c + d \<longrightarrow> norm (integral {a..c} f - integral {a..t} f) < e"
 | 
| 56188 | 4035 | proof - | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4036 |   have intm: "(\<lambda>x. f (- x)) integrable_on {-b .. -a}" "- b < - c" "- c \<le> - a"
 | 
| 53634 | 4037 | using assms by auto | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4038 | from indefinite_integral_continuous_left[OF intm \<open>e>0\<close>] | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4039 | obtain d where "0 < d" | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4040 | and d: "\<And>t. \<lbrakk>- c - d < t; t \<le> -c\<rbrakk> | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4041 |              \<Longrightarrow> norm (integral {- b..- c} (\<lambda>x. f (-x)) - integral {- b..t} (\<lambda>x. f (-x))) < e"
 | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4042 | by metis | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4043 | let ?d = "min d (b - c)" | 
| 53634 | 4044 | show ?thesis | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4045 | proof (intro that[of "?d"] allI impI) | 
| 53634 | 4046 | show "0 < ?d" | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4047 | using \<open>0 < d\<close> \<open>c < b\<close> by auto | 
| 53634 | 4048 | fix t :: real | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4049 | assume t: "c \<le> t \<and> t < c + ?d" | 
| 66402 | 4050 |     have *: "integral {a..c} f = integral {a..b} f - integral {c..b} f"
 | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4051 |             "integral {a..t} f = integral {a..b} f - integral {t..b} f"
 | 
| 63170 | 4052 | apply (simp_all only: algebra_simps) | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4053 | using assms t by (auto simp: integral_combine) | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4054 | have "(- c) - d < (- t)" "- t \<le> - c" | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4055 | using t by auto | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4056 |     from d[OF this] show "norm (integral {a..c} f - integral {a..t} f) < e"
 | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4057 | by (auto simp add: algebra_simps norm_minus_commute *) | 
| 53634 | 4058 | qed | 
| 4059 | qed | |
| 4060 | ||
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4061 | lemma indefinite_integral_continuous_1: | 
| 53634 | 4062 | fixes f :: "real \<Rightarrow> 'a::banach" | 
| 66402 | 4063 |   assumes "f integrable_on {a..b}"
 | 
| 4064 |   shows "continuous_on {a..b} (\<lambda>x. integral {a..x} f)"
 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4065 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4066 |   have "\<exists>d>0. \<forall>x'\<in>{a..b}. dist x' x < d \<longrightarrow> dist (integral {a..x'} f) (integral {a..x} f) < e" 
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4067 |        if x: "x \<in> {a..b}" and "e > 0" for x e :: real
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4068 | proof (cases "a = b") | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4069 | case True | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4070 | with that show ?thesis by force | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4071 | next | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4072 | case False | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4073 | with x have "a < b" by force | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4074 | with x consider "x = a" | "x = b" | "a < x" "x < b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4075 | by force | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4076 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4077 | proof cases | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4078 | case 1 show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4079 | apply (rule indefinite_integral_continuous_right [OF assms _ \<open>a < b\<close> \<open>e > 0\<close>], force) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4080 | using \<open>x = a\<close> apply (force simp: dist_norm algebra_simps) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4081 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4082 | next | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4083 | case 2 show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4084 | apply (rule indefinite_integral_continuous_left [OF assms \<open>a < b\<close> _ \<open>e > 0\<close>], force) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4085 | using \<open>x = b\<close> apply (force simp: dist_norm norm_minus_commute algebra_simps) | 
| 53634 | 4086 | done | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4087 | next | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4088 | case 3 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4089 | obtain d1 where "0 < d1" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4090 |         and d1: "\<And>t. \<lbrakk>x - d1 < t; t \<le> x\<rbrakk> \<Longrightarrow> norm (integral {a..x} f - integral {a..t} f) < e"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4091 | using 3 by (auto intro: indefinite_integral_continuous_left [OF assms \<open>a < x\<close> _ \<open>e > 0\<close>]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4092 | obtain d2 where "0 < d2" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4093 |         and d2: "\<And>t. \<lbrakk>x \<le> t; t < x + d2\<rbrakk> \<Longrightarrow> norm (integral {a..x} f - integral {a..t} f) < e"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4094 | using 3 by (auto intro: indefinite_integral_continuous_right [OF assms _ \<open>x < b\<close> \<open>e > 0\<close>]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4095 | show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4096 | proof (intro exI ballI conjI impI) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4097 | show "0 < min d1 d2" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4098 | using \<open>0 < d1\<close> \<open>0 < d2\<close> by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4099 |         show "dist (integral {a..y} f) (integral {a..x} f) < e"
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4100 |              if "y \<in> {a..b}" "dist y x < min d1 d2" for y
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4101 | proof (cases "y < x") | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4102 | case True | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4103 | with that d1 show ?thesis by (auto simp: dist_commute dist_norm) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4104 | next | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4105 | case False | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4106 | with that d2 show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4107 | by (auto simp: dist_commute dist_norm) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4108 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4109 | qed | 
| 53634 | 4110 | qed | 
| 4111 | qed | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4112 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4113 | by (auto simp: continuous_on_iff) | 
| 53634 | 4114 | qed | 
| 4115 | ||
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4116 | lemma indefinite_integral_continuous_1': | 
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 4117 | fixes f::"real \<Rightarrow> 'a::banach" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 4118 |   assumes "f integrable_on {a..b}"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 4119 |   shows "continuous_on {a..b} (\<lambda>x. integral {x..b} f)"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 4120 | proof - | 
| 66402 | 4121 |   have "integral {a..b} f - integral {a..x} f = integral {x..b} f" if "x \<in> {a..b}" for x
 | 
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 4122 | using integral_combine[OF _ _ assms, of x] that | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 4123 | by (auto simp: algebra_simps) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 4124 | with _ show ?thesis | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4125 | by (rule continuous_on_eq) (auto intro!: continuous_intros indefinite_integral_continuous_1 assms) | 
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 4126 | qed | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 4127 | |
| 35751 | 4128 | |
| 60420 | 4129 | subsection \<open>This doesn't directly involve integration, but that gives an easy proof.\<close> | 
| 35751 | 4130 | |
| 53634 | 4131 | lemma has_derivative_zero_unique_strong_interval: | 
| 4132 | fixes f :: "real \<Rightarrow> 'a::banach" | |
| 4133 | assumes "finite k" | |
| 66402 | 4134 |     and "continuous_on {a..b} f"
 | 
| 53634 | 4135 | and "f a = y" | 
| 66402 | 4136 |     and "\<forall>x\<in>({a..b} - k). (f has_derivative (\<lambda>h. 0)) (at x within {a..b})" "x \<in> {a..b}"
 | 
| 35751 | 4137 | shows "f x = y" | 
| 53634 | 4138 | proof - | 
| 4139 | have ab: "a \<le> b" | |
| 4140 | using assms by auto | |
| 4141 | have *: "a \<le> x" | |
| 4142 | using assms(5) by auto | |
| 66402 | 4143 |   have "((\<lambda>x. 0::'a) has_integral f x - f a) {a..x}"
 | 
| 53634 | 4144 | apply (rule fundamental_theorem_of_calculus_interior_strong[OF assms(1) *]) | 
| 4145 | apply (rule continuous_on_subset[OF assms(2)]) | |
| 4146 | defer | |
| 4147 | apply safe | |
| 4148 | unfolding has_vector_derivative_def | |
| 4149 | apply (subst has_derivative_within_open[symmetric]) | |
| 4150 | apply assumption | |
| 56188 | 4151 | apply (rule open_greaterThanLessThan) | 
| 66402 | 4152 |     apply (rule has_derivative_within_subset[where s="{a..b}"])
 | 
| 53634 | 4153 | using assms(4) assms(5) | 
| 56188 | 4154 | apply (auto simp: mem_box) | 
| 53634 | 4155 | done | 
| 4156 | note this[unfolded *] | |
| 35751 | 4157 | note has_integral_unique[OF has_integral_0 this] | 
| 53634 | 4158 | then show ?thesis | 
| 4159 | unfolding assms by auto | |
| 4160 | qed | |
| 4161 | ||
| 35751 | 4162 | |
| 60420 | 4163 | subsection \<open>Generalize a bit to any convex set.\<close> | 
| 35751 | 4164 | |
| 53634 | 4165 | lemma has_derivative_zero_unique_strong_convex: | 
| 56188 | 4166 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4167 | assumes "convex S" "finite K" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4168 | and contf: "continuous_on S f" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4169 | and "c \<in> S" "f c = y" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4170 | and derf: "\<And>x. x \<in> (S - K) \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within S)" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4171 | and "x \<in> S" | 
| 35751 | 4172 | shows "f x = y" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4173 | proof (cases "x = c") | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4174 | case True with \<open>f c = y\<close> show ?thesis | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4175 | by blast | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4176 | next | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4177 | case False | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4178 | let ?\<phi> = "\<lambda>u. (1 - u) *\<^sub>R c + u *\<^sub>R x" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4179 |   have contf': "continuous_on {0 ..1} (f \<circ> ?\<phi>)"
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4180 | apply (rule continuous_intros continuous_on_subset[OF contf])+ | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4181 | using \<open>convex S\<close> \<open>x \<in> S\<close> \<open>c \<in> S\<close> by (auto simp add: convex_alt algebra_simps) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4182 | have "t = u" if "?\<phi> t = ?\<phi> u" for t u | 
| 53634 | 4183 | proof - | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4184 | from that have "(t - u) *\<^sub>R x = (t - u) *\<^sub>R c" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4185 | by (auto simp add: algebra_simps) | 
| 61165 | 4186 | then show ?thesis | 
| 60420 | 4187 | using \<open>x \<noteq> c\<close> by auto | 
| 53634 | 4188 | qed | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4189 | then have eq: "(SOME t. ?\<phi> t = ?\<phi> u) = u" for u | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4190 | by blast | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4191 | then have "(?\<phi> -` K) \<subseteq> (\<lambda>z. SOME t. ?\<phi> t = z) ` K" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4192 | by (clarsimp simp: image_iff) (metis (no_types) eq) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4193 | then have fin: "finite (?\<phi> -` K)" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4194 | by (rule finite_surj[OF \<open>finite K\<close>]) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4195 | |
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4196 |   have derf': "((\<lambda>u. f (?\<phi> u)) has_derivative (\<lambda>h. 0)) (at t within {0..1})"
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4197 |                if "t \<in> {0..1} - {t. ?\<phi> t \<in> K}" for t
 | 
| 53634 | 4198 | proof - | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4199 |     have df: "(f has_derivative (\<lambda>h. 0)) (at (?\<phi> t) within ?\<phi> ` {0..1})"
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4200 | apply (rule has_derivative_within_subset [OF derf]) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4201 | using \<open>convex S\<close> \<open>x \<in> S\<close> \<open>c \<in> S\<close> that by (auto simp add: convex_alt algebra_simps) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4202 |     have "(f \<circ> ?\<phi> has_derivative (\<lambda>x. 0) \<circ> (\<lambda>z. (0 - z *\<^sub>R c) + z *\<^sub>R x)) (at t within {0..1})"
 | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4203 | by (rule derivative_eq_intros df | simp)+ | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4204 | then show ?thesis | 
| 53634 | 4205 | unfolding o_def . | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4206 | qed | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4207 | have "(f \<circ> ?\<phi>) 1 = y" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4208 | apply (rule has_derivative_zero_unique_strong_interval[OF fin contf']) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4209 | unfolding o_def using \<open>f c = y\<close> derf' by auto | 
| 53634 | 4210 | then show ?thesis | 
| 4211 | by auto | |
| 4212 | qed | |
| 4213 | ||
| 4214 | ||
| 60420 | 4215 | text \<open>Also to any open connected set with finite set of exceptions. Could | 
| 4216 | generalize to locally convex set with limpt-free set of exceptions.\<close> | |
| 35751 | 4217 | |
| 53634 | 4218 | lemma has_derivative_zero_unique_strong_connected: | 
| 56188 | 4219 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4220 | assumes "connected S" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4221 | and "open S" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4222 | and "finite K" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4223 | and contf: "continuous_on S f" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4224 | and "c \<in> S" | 
| 53634 | 4225 | and "f c = y" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4226 | and derf: "\<And>x. x \<in> (S - K) \<Longrightarrow> (f has_derivative (\<lambda>h. 0)) (at x within S)" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4227 | and "x \<in> S" | 
| 35751 | 4228 | shows "f x = y" | 
| 53634 | 4229 | proof - | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 4230 |   have "\<exists>e>0. ball x e \<subseteq> (S \<inter> f -` {f x})" if "x \<in> S" for x
 | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4231 | proof - | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4232 | obtain e where "0 < e" and e: "ball x e \<subseteq> S" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4233 | using \<open>x \<in> S\<close> \<open>open S\<close> open_contains_ball by blast | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4234 |     have "ball x e \<subseteq> {u \<in> S. f u \<in> {f x}}"
 | 
| 53634 | 4235 | proof safe | 
| 4236 | fix y | |
| 4237 | assume y: "y \<in> ball x e" | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4238 | then show "y \<in> S" | 
| 53634 | 4239 | using e by auto | 
| 4240 | show "f y = f x" | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4241 | proof (rule has_derivative_zero_unique_strong_convex[OF convex_ball \<open>finite K\<close>]) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4242 | show "continuous_on (ball x e) f" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4243 | using contf continuous_on_subset e by blast | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4244 | show "(f has_derivative (\<lambda>h. 0)) (at u within ball x e)" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4245 | if "u \<in> ball x e - K" for u | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4246 | by (metis Diff_iff contra_subsetD derf e has_derivative_within_subset that) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4247 | qed (use y e \<open>0 < e\<close> in auto) | 
| 53634 | 4248 | qed | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 4249 |     then show "\<exists>e>0. ball x e \<subseteq> (S \<inter> f -` {f x})"
 | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4250 | using \<open>0 < e\<close> by blast | 
| 53634 | 4251 | qed | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 4252 |   then have "openin (subtopology euclidean S) (S \<inter> f -` {y})"
 | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4253 | by (auto intro!: open_openin_trans[OF \<open>open S\<close>] simp: open_contains_ball) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 4254 |   moreover have "closedin (subtopology euclidean S) (S \<inter> f -` {y})"
 | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4255 | by (force intro!: continuous_closedin_preimage [OF contf]) | 
| 66884 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 4256 |   ultimately have "(S \<inter> f -` {y}) = {} \<or> (S \<inter> f -` {y}) = S"
 | 
| 
c2128ab11f61
Switching to inverse image and constant_on, plus some new material
 paulson <lp15@cam.ac.uk> parents: 
66793diff
changeset | 4257 | using \<open>connected S\<close> by (simp add: connected_clopen) | 
| 53634 | 4258 | then show ?thesis | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4259 | using \<open>x \<in> S\<close> \<open>f c = y\<close> \<open>c \<in> S\<close> by auto | 
| 53634 | 4260 | qed | 
| 4261 | ||
| 56332 | 4262 | lemma has_derivative_zero_connected_constant: | 
| 4263 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4264 | assumes "connected S" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4265 | and "open S" | 
| 56332 | 4266 | and "finite k" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4267 | and "continuous_on S f" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4268 | and "\<forall>x\<in>(S - k). (f has_derivative (\<lambda>h. 0)) (at x within S)" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4269 | obtains c where "\<And>x. x \<in> S \<Longrightarrow> f(x) = c" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4270 | proof (cases "S = {}")
 | 
| 56332 | 4271 | case True | 
| 4272 | then show ?thesis | |
| 4273 | by (metis empty_iff that) | |
| 4274 | next | |
| 4275 | case False | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 4276 | then obtain c where "c \<in> S" | 
| 56332 | 4277 | by (metis equals0I) | 
| 4278 | then show ?thesis | |
| 4279 | by (metis has_derivative_zero_unique_strong_connected assms that) | |
| 4280 | qed | |
| 4281 | ||
| 53634 | 4282 | |
| 60420 | 4283 | subsection \<open>Integrating characteristic function of an interval\<close> | 
| 53634 | 4284 | |
| 4285 | lemma has_integral_restrict_open_subinterval: | |
| 56188 | 4286 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4287 | assumes intf: "(f has_integral i) (cbox c d)" | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4288 | and cb: "cbox c d \<subseteq> cbox a b" | 
| 56188 | 4289 | shows "((\<lambda>x. if x \<in> box c d then f x else 0) has_integral i) (cbox a b)" | 
| 66535 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4290 | proof (cases "cbox c d = {}")
 | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4291 | case True | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4292 |   then have "box c d = {}"
 | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4293 | by (metis bot.extremum_uniqueI box_subset_cbox) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4294 | then show ?thesis | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4295 | using True intf by auto | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4296 | next | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4297 | case False | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4298 | then obtain p where pdiv: "p division_of cbox a b" and inp: "cbox c d \<in> p" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4299 | using cb partial_division_extend_1 by blast | 
| 63040 | 4300 | define g where [abs_def]: "g x = (if x \<in>box c d then f x else 0)" for x | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 4301 | interpret operative "lift_option plus" "Some (0 :: 'b)" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 4302 | "\<lambda>i. if g integrable_on i then Some (integral i g) else None" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 4303 | by (fact operative_integralI) | 
| 66535 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4304 | note operat = division [OF pdiv, symmetric] | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4305 | show ?thesis | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4306 | proof (cases "(g has_integral i) (cbox a b)") | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4307 | case True then show ?thesis | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4308 | by (simp add: g_def) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4309 | next | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4310 | case False | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4311 |     have iterate:"F (\<lambda>i. if g integrable_on i then Some (integral i g) else None) (p - {cbox c d}) = Some 0"
 | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4312 | proof (intro neutral ballI) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4313 | fix x | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4314 |       assume x: "x \<in> p - {cbox c d}"
 | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4315 | then have "x \<in> p" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4316 | by auto | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4317 | then obtain u v where uv: "x = cbox u v" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4318 | using pdiv by blast | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4319 |       have "interior x \<inter> interior (cbox c d) = {}"
 | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4320 | using pdiv inp x by blast | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4321 | then have "(g has_integral 0) x" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4322 | unfolding uv using has_integral_spike_interior[where f="\<lambda>x. 0"] | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4323 | by (metis (no_types, lifting) disjoint_iff_not_equal g_def has_integral_0_eq interior_cbox) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4324 | then show "(if g integrable_on x then Some (integral x g) else None) = Some 0" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4325 | by auto | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4326 | qed | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4327 | interpret comm_monoid_set "lift_option plus" "Some (0 :: 'b)" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4328 | by (intro comm_monoid_set.intro comm_monoid_lift_option add.comm_monoid_axioms) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4329 | have intg: "g integrable_on cbox c d" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4330 | using integrable_spike_interior[where f=f] | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4331 | by (meson g_def has_integral_integrable intf) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4332 | moreover have "integral (cbox c d) g = i" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4333 | proof (rule has_integral_unique[OF has_integral_spike_interior intf]) | 
| 66539 
0ad3fc48c9ec
final cleanup of negligible_standard_hyperplane and other things
 paulson <lp15@cam.ac.uk> parents: 
66537diff
changeset | 4334 | show "\<And>x. x \<in> box c d \<Longrightarrow> f x = g x" | 
| 66535 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4335 | by (auto simp: g_def) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4336 | show "(g has_integral integral (cbox c d) g) (cbox c d)" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4337 | by (rule integrable_integral[OF intg]) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4338 | qed | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4339 | ultimately have "F (\<lambda>A. if g integrable_on A then Some (integral A g) else None) p = Some i" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4340 | by (metis (full_types, lifting) division_of_finite inp iterate pdiv remove right_neutral) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4341 | then | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4342 | have "(g has_integral i) (cbox a b)" | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4343 | by (metis integrable_on_def integral_unique operat option.inject option.simps(3)) | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4344 | with False show ?thesis | 
| 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4345 | by blast | 
| 35751 | 4346 | qed | 
| 53634 | 4347 | qed | 
| 4348 | ||
| 66535 
64035d9161d3
unscrambled has_integral_restrict_open_subinterval
 paulson <lp15@cam.ac.uk> parents: 
66533diff
changeset | 4349 | |
| 53634 | 4350 | lemma has_integral_restrict_closed_subinterval: | 
| 56188 | 4351 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | 
| 4352 | assumes "(f has_integral i) (cbox c d)" | |
| 4353 | and "cbox c d \<subseteq> cbox a b" | |
| 4354 | shows "((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b)" | |
| 53634 | 4355 | proof - | 
| 4356 | note has_integral_restrict_open_subinterval[OF assms] | |
| 35751 | 4357 | note * = has_integral_spike[OF negligible_frontier_interval _ this] | 
| 53634 | 4358 | show ?thesis | 
| 4359 | apply (rule *[of c d]) | |
| 56188 | 4360 | using box_subset_cbox[of c d] | 
| 53634 | 4361 | apply auto | 
| 4362 | done | |
| 4363 | qed | |
| 4364 | ||
| 4365 | lemma has_integral_restrict_closed_subintervals_eq: | |
| 56188 | 4366 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b::banach" | 
| 4367 | assumes "cbox c d \<subseteq> cbox a b" | |
| 4368 | shows "((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b) \<longleftrightarrow> (f has_integral i) (cbox c d)" | |
| 53634 | 4369 | (is "?l = ?r") | 
| 56188 | 4370 | proof (cases "cbox c d = {}")
 | 
| 53634 | 4371 | case False | 
| 56188 | 4372 | let ?g = "\<lambda>x. if x \<in> cbox c d then f x else 0" | 
| 53634 | 4373 | show ?thesis | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 4374 | proof | 
| 53634 | 4375 | assume ?l | 
| 56188 | 4376 | then have "?g integrable_on cbox c d" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 4377 | using assms has_integral_integrable integrable_subinterval by blast | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 4378 | then have "f integrable_on cbox c d" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 4379 | by (rule integrable_eq) auto | 
| 66536 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 4380 | moreover then have "i = integral (cbox c d) f" | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 4381 | by (meson \<open>((\<lambda>x. if x \<in> cbox c d then f x else 0) has_integral i) (cbox a b)\<close> assms has_integral_restrict_closed_subinterval has_integral_unique integrable_integral) | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 4382 | ultimately show ?r by auto | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 4383 | next | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 4384 | assume ?r then show ?l | 
| 
9c95b2b54337
Unscrambling continues as far as negligible_standard_hyperplane
 paulson <lp15@cam.ac.uk> parents: 
66535diff
changeset | 4385 | by (rule has_integral_restrict_closed_subinterval[OF _ assms]) | 
| 53634 | 4386 | qed | 
| 4387 | qed auto | |
| 4388 | ||
| 4389 | ||
| 60420 | 4390 | text \<open>Hence we can apply the limit process uniformly to all integrals.\<close> | 
| 53634 | 4391 | |
| 4392 | lemma has_integral': | |
| 56188 | 4393 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66504 | 4394 | shows "(f has_integral i) S \<longleftrightarrow> | 
| 56188 | 4395 | (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | 
| 66504 | 4396 | (\<exists>z. ((\<lambda>x. if x \<in> S then f(x) else 0) has_integral z) (cbox a b) \<and> norm(z - i) < e))" | 
| 53634 | 4397 | (is "?l \<longleftrightarrow> (\<forall>e>0. ?r e)") | 
| 66504 | 4398 | proof (cases "\<exists>a b. S = cbox a b") | 
| 4399 | case False then show ?thesis | |
| 4400 | by (simp add: has_integral_alt) | |
| 4401 | next | |
| 4402 | case True | |
| 4403 | then obtain a b where S: "S = cbox a b" | |
| 4404 | by blast | |
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4405 | obtain B where " 0 < B" and B: "\<And>x. x \<in> cbox a b \<Longrightarrow> norm x \<le> B" | 
| 66504 | 4406 | using bounded_cbox[unfolded bounded_pos] by blast | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4407 | show ?thesis | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4408 | proof safe | 
| 53634 | 4409 | fix e :: real | 
| 4410 | assume ?l and "e > 0" | |
| 66504 | 4411 | have "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) (cbox c d)" | 
| 4412 | if "ball 0 (B+1) \<subseteq> cbox c d" for c d | |
| 4413 | unfolding S using B that | |
| 4414 | by (force intro: \<open>?l\<close>[unfolded S] has_integral_restrict_closed_subinterval) | |
| 4415 | then show "?r e" | |
| 53634 | 4416 | apply (rule_tac x="B+1" in exI) | 
| 66504 | 4417 | using \<open>B>0\<close> \<open>e>0\<close> by force | 
| 53634 | 4418 | next | 
| 4419 | assume as: "\<forall>e>0. ?r e" | |
| 66504 | 4420 | then obtain C | 
| 4421 | where C: "\<And>a b. ball 0 C \<subseteq> cbox a b \<Longrightarrow> | |
| 4422 | \<exists>z. ((\<lambda>x. if x \<in> S then f x else 0) has_integral z) (cbox a b)" | |
| 4423 | by (meson zero_less_one) | |
| 63040 | 4424 | define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)" | 
| 4425 | define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)" | |
| 66504 | 4426 | have "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" if "norm x \<le> B" "i \<in> Basis" for x i | 
| 4427 | using that and Basis_le_norm[OF \<open>i\<in>Basis\<close>, of x] | |
| 4428 | by (auto simp add: field_simps sum_negf c_def d_def) | |
| 4429 | then have c_d: "cbox a b \<subseteq> cbox c d" | |
| 4430 | by (meson B mem_box(2) subsetI) | |
| 4431 | have "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" | |
| 4432 | if x: "norm (0 - x) < C" and i: "i \<in> Basis" for x i | |
| 4433 | using Basis_le_norm[OF i, of x] x i by (auto simp: sum_negf c_def d_def) | |
| 4434 | then have "ball 0 C \<subseteq> cbox c d" | |
| 4435 | by (auto simp: mem_box dist_norm) | |
| 4436 | with C obtain y where y: "(f has_integral y) (cbox a b)" | |
| 4437 | using c_d has_integral_restrict_closed_subintervals_eq S by blast | |
| 53634 | 4438 | have "y = i" | 
| 4439 | proof (rule ccontr) | |
| 66504 | 4440 | assume "y \<noteq> i" | 
| 53634 | 4441 | then have "0 < norm (y - i)" | 
| 4442 | by auto | |
| 66504 | 4443 | from as[rule_format,OF this] | 
| 4444 | obtain C where C: "\<And>a b. ball 0 C \<subseteq> cbox a b \<Longrightarrow> | |
| 4445 | \<exists>z. ((\<lambda>x. if x \<in> S then f x else 0) has_integral z) (cbox a b) \<and> norm (z-i) < norm (y-i)" | |
| 4446 | by auto | |
| 63040 | 4447 | define c :: 'n where "c = (\<Sum>i\<in>Basis. (- max B C) *\<^sub>R i)" | 
| 4448 | define d :: 'n where "d = (\<Sum>i\<in>Basis. max B C *\<^sub>R i)" | |
| 66504 | 4449 | have "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" | 
| 4450 | if "norm x \<le> B" and "i \<in> Basis" for x i | |
| 4451 | using that Basis_le_norm[of i x] by (auto simp add: field_simps sum_negf c_def d_def) | |
| 4452 | then have c_d: "cbox a b \<subseteq> cbox c d" | |
| 4453 | by (simp add: B mem_box(2) subset_eq) | |
| 4454 | have "c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" if "norm (0 - x) < C" and "i \<in> Basis" for x i | |
| 4455 | using Basis_le_norm[of i x] that by (auto simp: sum_negf c_def d_def) | |
| 4456 | then have "ball 0 C \<subseteq> cbox c d" | |
| 4457 | by (auto simp: mem_box dist_norm) | |
| 4458 | with C obtain z where z: "(f has_integral z) (cbox a b)" "norm (z-i) < norm (y-i)" | |
| 4459 | using has_integral_restrict_closed_subintervals_eq[OF c_d] S by blast | |
| 4460 | moreover then have "z = y" | |
| 4461 | by (blast intro: has_integral_unique[OF _ y]) | |
| 4462 | ultimately show False | |
| 53634 | 4463 | by auto | 
| 4464 | qed | |
| 4465 | then show ?l | |
| 66504 | 4466 | using y by (auto simp: S) | 
| 53634 | 4467 | qed | 
| 4468 | qed | |
| 4469 | ||
| 4470 | lemma has_integral_le: | |
| 56188 | 4471 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 66504 | 4472 | assumes fg: "(f has_integral i) S" "(g has_integral j) S" | 
| 4473 | and le: "\<And>x. x \<in> S \<Longrightarrow> f x \<le> g x" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 4474 | shows "i \<le> j" | 
| 66504 | 4475 | using has_integral_component_le[OF _ fg, of 1] le by auto | 
| 53634 | 4476 | |
| 4477 | lemma integral_le: | |
| 56188 | 4478 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 4479 | assumes "f integrable_on S" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 4480 | and "g integrable_on S" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 4481 | and "\<And>x. x \<in> S \<Longrightarrow> f x \<le> g x" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 4482 | shows "integral S f \<le> integral S g" | 
| 53634 | 4483 | by (rule has_integral_le[OF assms(1,2)[unfolded has_integral_integral] assms(3)]) | 
| 4484 | ||
| 4485 | lemma has_integral_nonneg: | |
| 56188 | 4486 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 4487 | assumes "(f has_integral i) S" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 4488 | and "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x" | 
| 53634 | 4489 | shows "0 \<le> i" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 4490 | using has_integral_component_nonneg[of 1 f i S] | 
| 53634 | 4491 | unfolding o_def | 
| 4492 | using assms | |
| 4493 | by auto | |
| 4494 | ||
| 4495 | lemma integral_nonneg: | |
| 56188 | 4496 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 4497 | assumes f: "f integrable_on S" and 0: "\<And>x. x \<in> S \<Longrightarrow> 0 \<le> f x" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 4498 | shows "0 \<le> integral S f" | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 4499 | by (rule has_integral_nonneg[OF f[unfolded has_integral_integral] 0]) | 
| 53634 | 4500 | |
| 4501 | ||
| 60420 | 4502 | text \<open>Hence a general restriction property.\<close> | 
| 53634 | 4503 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4504 | lemma has_integral_restrict [simp]: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4505 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4506 | assumes "S \<subseteq> T" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4507 | shows "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) T \<longleftrightarrow> (f has_integral i) S" | 
| 53634 | 4508 | proof - | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4509 | have *: "\<And>x. (if x \<in> T then if x \<in> S then f x else 0 else 0) = (if x\<in>S then f x else 0)" | 
| 53634 | 4510 | using assms by auto | 
| 4511 | show ?thesis | |
| 4512 | apply (subst(2) has_integral') | |
| 4513 | apply (subst has_integral') | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4514 | apply (simp add: *) | 
| 53634 | 4515 | done | 
| 4516 | qed | |
| 4517 | ||
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4518 | corollary has_integral_restrict_UNIV: | 
| 56188 | 4519 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 53634 | 4520 | shows "((\<lambda>x. if x \<in> s then f x else 0) has_integral i) UNIV \<longleftrightarrow> (f has_integral i) s" | 
| 4521 | by auto | |
| 4522 | ||
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4523 | lemma has_integral_restrict_Int: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4524 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4525 | shows "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) T \<longleftrightarrow> (f has_integral i) (S \<inter> T)" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4526 | proof - | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4527 | have "((\<lambda>x. if x \<in> T then if x \<in> S then f x else 0 else 0) has_integral i) UNIV = | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4528 | ((\<lambda>x. if x \<in> S \<inter> T then f x else 0) has_integral i) UNIV" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4529 | by (rule has_integral_cong) auto | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4530 | then show ?thesis | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4531 | using has_integral_restrict_UNIV by fastforce | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4532 | qed | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4533 | |
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4534 | lemma integral_restrict_Int: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4535 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4536 | shows "integral T (\<lambda>x. if x \<in> S then f x else 0) = integral (S \<inter> T) f" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4537 | by (metis (no_types, lifting) has_integral_cong has_integral_restrict_Int integrable_integral integral_unique not_integrable_integral) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4538 | |
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4539 | lemma integrable_restrict_Int: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4540 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4541 | shows "(\<lambda>x. if x \<in> S then f x else 0) integrable_on T \<longleftrightarrow> f integrable_on (S \<inter> T)" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4542 | using has_integral_restrict_Int by fastforce | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4543 | |
| 53634 | 4544 | lemma has_integral_on_superset: | 
| 56188 | 4545 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4546 | assumes f: "(f has_integral i) S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4547 | and "\<And>x. x \<notin> S \<Longrightarrow> f x = 0" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4548 | and "S \<subseteq> T" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4549 | shows "(f has_integral i) T" | 
| 53634 | 4550 | proof - | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4551 | have "(\<lambda>x. if x \<in> S then f x else 0) = (\<lambda>x. if x \<in> T then f x else 0)" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4552 | using assms by fastforce | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4553 | with f show ?thesis | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4554 | by (simp only: has_integral_restrict_UNIV [symmetric, of f]) | 
| 53634 | 4555 | qed | 
| 4556 | ||
| 4557 | lemma integrable_on_superset: | |
| 56188 | 4558 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 4559 | assumes "f integrable_on S" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 4560 | and "\<And>x. x \<notin> S \<Longrightarrow> f x = 0" | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 4561 | and "S \<subseteq> t" | 
| 35751 | 4562 | shows "f integrable_on t" | 
| 53634 | 4563 | using assms | 
| 4564 | unfolding integrable_on_def | |
| 4565 | by (auto intro:has_integral_on_superset) | |
| 4566 | ||
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4567 | lemma integral_restrict_UNIV [intro]: | 
| 56188 | 4568 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 4569 | shows "f integrable_on S \<Longrightarrow> integral UNIV (\<lambda>x. if x \<in> S then f x else 0) = integral S f" | 
| 53634 | 4570 | apply (rule integral_unique) | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 4571 | unfolding has_integral_restrict_UNIV | 
| 53634 | 4572 | apply auto | 
| 4573 | done | |
| 4574 | ||
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 4575 | lemma integrable_restrict_UNIV: | 
| 56188 | 4576 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 53634 | 4577 | shows "(\<lambda>x. if x \<in> s then f x else 0) integrable_on UNIV \<longleftrightarrow> f integrable_on s" | 
| 4578 | unfolding integrable_on_def | |
| 4579 | by auto | |
| 4580 | ||
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4581 | lemma has_integral_subset_component_le: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4582 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4583 | assumes k: "k \<in> Basis" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4584 | and as: "S \<subseteq> T" "(f has_integral i) S" "(f has_integral j) T" "\<And>x. x\<in>T \<Longrightarrow> 0 \<le> f(x)\<bullet>k" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4585 | shows "i\<bullet>k \<le> j\<bullet>k" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4586 | proof - | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4587 | have "((\<lambda>x. if x \<in> S then f x else 0) has_integral i) UNIV" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4588 | "((\<lambda>x. if x \<in> T then f x else 0) has_integral j) UNIV" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4589 | by (simp_all add: assms) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4590 | then show ?thesis | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4591 | apply (rule has_integral_component_le[OF k]) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4592 | using as by auto | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4593 | qed | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4594 | |
| 56188 | 4595 | lemma negligible_on_intervals: "negligible s \<longleftrightarrow> (\<forall>a b. negligible(s \<inter> cbox a b))" (is "?l \<longleftrightarrow> ?r") | 
| 53634 | 4596 | proof | 
| 4597 | assume ?r | |
| 4598 | show ?l | |
| 4599 | unfolding negligible_def | |
| 4600 | proof safe | |
| 61165 | 4601 | fix a b | 
| 4602 | show "(indicator s has_integral 0) (cbox a b)" | |
| 60420 | 4603 | apply (rule has_integral_negligible[OF \<open>?r\<close>[rule_format,of a b]]) | 
| 53634 | 4604 | unfolding indicator_def | 
| 4605 | apply auto | |
| 4606 | done | |
| 4607 | qed | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 4608 | qed (simp add: negligible_Int) | 
| 53634 | 4609 | |
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4610 | lemma negligible_translation: | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4611 | assumes "negligible S" | 
| 67399 | 4612 | shows "negligible ((+) c ` S)" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4613 | proof - | 
| 67399 | 4614 | have inj: "inj ((+) c)" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4615 | by simp | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4616 | show ?thesis | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4617 | using assms | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4618 | proof (clarsimp simp: negligible_def) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4619 | fix a b | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4620 | assume "\<forall>x y. (indicator S has_integral 0) (cbox x y)" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4621 | then have *: "(indicator S has_integral 0) (cbox (a-c) (b-c))" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4622 | by (meson Diff_iff assms has_integral_negligible indicator_simps(2)) | 
| 67399 | 4623 | have eq: "indicator ((+) c ` S) = (\<lambda>x. indicator S (x - c))" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4624 | by (force simp add: indicator_def) | 
| 67399 | 4625 | show "(indicator ((+) c ` S) has_integral 0) (cbox a b)" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4626 | using has_integral_affinity [OF *, of 1 "-c"] | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4627 | cbox_translation [of "c" "-c+a" "-c+b"] | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4628 | by (simp add: eq add.commute) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4629 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4630 | qed | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4631 | |
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4632 | lemma negligible_translation_rev: | 
| 67399 | 4633 | assumes "negligible ((+) c ` S)" | 
| 63492 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4634 | shows "negligible S" | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4635 | by (metis negligible_translation [OF assms, of "-c"] translation_galois) | 
| 
a662e8139804
More advanced theorems about retracts, homotopies., etc
 paulson <lp15@cam.ac.uk> parents: 
63469diff
changeset | 4636 | |
| 53634 | 4637 | lemma has_integral_spike_set_eq: | 
| 56188 | 4638 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4639 | assumes "negligible ((S - T) \<union> (T - S))" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4640 | shows "(f has_integral y) S \<longleftrightarrow> (f has_integral y) T" | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 4641 | unfolding has_integral_restrict_UNIV[symmetric,of f] | 
| 53634 | 4642 | apply (rule has_integral_spike_eq[OF assms]) | 
| 62390 | 4643 | by (auto split: if_split_asm) | 
| 53634 | 4644 | |
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 4645 | lemma has_integral_spike_set: | 
| 56188 | 4646 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4647 | assumes "(f has_integral y) S" "negligible ((S - T) \<union> (T - S))" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4648 | shows "(f has_integral y) T" | 
| 53634 | 4649 | using assms has_integral_spike_set_eq | 
| 4650 | by auto | |
| 4651 | ||
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 4652 | lemma integrable_spike_set: | 
| 56188 | 4653 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4654 | assumes "f integrable_on S" and "negligible ((S - T) \<union> (T - S))" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4655 | shows "f integrable_on T" | 
| 63945 
444eafb6e864
a few new theorems and a renaming
 paulson <lp15@cam.ac.uk> parents: 
63944diff
changeset | 4656 | using assms by (simp add: integrable_on_def has_integral_spike_set_eq) | 
| 35751 | 4657 | |
| 53634 | 4658 | lemma integrable_spike_set_eq: | 
| 56188 | 4659 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4660 | assumes "negligible ((S - T) \<union> (T - S))" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4661 | shows "f integrable_on S \<longleftrightarrow> f integrable_on T" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4662 | by (blast intro: integrable_spike_set assms negligible_subset) | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4663 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4664 | lemma has_integral_interior: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4665 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4666 | shows "negligible(frontier S) \<Longrightarrow> (f has_integral y) (interior S) \<longleftrightarrow> (f has_integral y) S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4667 | apply (rule has_integral_spike_set_eq) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4668 | apply (auto simp: frontier_def Un_Diff closure_def) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4669 | apply (metis Diff_eq_empty_iff interior_subset negligible_empty) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4670 | done | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4671 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4672 | lemma has_integral_closure: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4673 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4674 | shows "negligible(frontier S) \<Longrightarrow> (f has_integral y) (closure S) \<longleftrightarrow> (f has_integral y) S" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4675 | apply (rule has_integral_spike_set_eq) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4676 | apply (auto simp: Un_Diff closure_Un_frontier negligible_diff) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4677 | by (simp add: Diff_eq closure_Un_frontier) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4678 | |
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4679 | lemma has_integral_open_interval: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4680 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4681 | shows "(f has_integral y) (box a b) \<longleftrightarrow> (f has_integral y) (cbox a b)" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4682 | unfolding interior_cbox [symmetric] | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4683 | by (metis frontier_cbox has_integral_interior negligible_frontier_interval) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4684 | |
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4685 | lemma integrable_on_open_interval: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4686 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4687 | shows "f integrable_on box a b \<longleftrightarrow> f integrable_on cbox a b" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4688 | by (simp add: has_integral_open_interval integrable_on_def) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4689 | |
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4690 | lemma integral_open_interval: | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4691 | fixes f :: "'a :: euclidean_space \<Rightarrow> 'b :: banach" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4692 | shows "integral(box a b) f = integral(cbox a b) f" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4693 | by (metis has_integral_integrable_integral has_integral_open_interval not_integrable_integral) | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4694 | |
| 35751 | 4695 | |
| 60420 | 4696 | subsection \<open>More lemmas that are useful later\<close> | 
| 53634 | 4697 | |
| 4698 | lemma has_integral_subset_le: | |
| 56188 | 4699 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 53634 | 4700 | assumes "s \<subseteq> t" | 
| 4701 | and "(f has_integral i) s" | |
| 4702 | and "(f has_integral j) t" | |
| 4703 | and "\<forall>x\<in>t. 0 \<le> f x" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 4704 | shows "i \<le> j" | 
| 53634 | 4705 | using has_integral_subset_component_le[OF _ assms(1), of 1 f i j] | 
| 4706 | using assms | |
| 4707 | by auto | |
| 4708 | ||
| 4709 | lemma integral_subset_component_le: | |
| 56188 | 4710 | fixes f :: "'n::euclidean_space \<Rightarrow> 'm::euclidean_space" | 
| 53634 | 4711 | assumes "k \<in> Basis" | 
| 4712 | and "s \<subseteq> t" | |
| 4713 | and "f integrable_on s" | |
| 4714 | and "f integrable_on t" | |
| 4715 | and "\<forall>x \<in> t. 0 \<le> f x \<bullet> k" | |
| 50526 
899c9c4e4a4c
Remove the indexed basis from the definition of euclidean spaces and only use the set of Basis vectors
 hoelzl parents: 
50348diff
changeset | 4716 | shows "(integral s f)\<bullet>k \<le> (integral t f)\<bullet>k" | 
| 53634 | 4717 | apply (rule has_integral_subset_component_le) | 
| 4718 | using assms | |
| 4719 | apply auto | |
| 4720 | done | |
| 4721 | ||
| 4722 | lemma integral_subset_le: | |
| 56188 | 4723 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 53634 | 4724 | assumes "s \<subseteq> t" | 
| 4725 | and "f integrable_on s" | |
| 4726 | and "f integrable_on t" | |
| 4727 | and "\<forall>x \<in> t. 0 \<le> f x" | |
| 4728 | shows "integral s f \<le> integral t f" | |
| 4729 | apply (rule has_integral_subset_le) | |
| 4730 | using assms | |
| 4731 | apply auto | |
| 4732 | done | |
| 4733 | ||
| 4734 | lemma has_integral_alt': | |
| 56188 | 4735 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4736 | shows "(f has_integral i) s \<longleftrightarrow> | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4737 | (\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b) \<and> | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4738 | (\<forall>e>0. \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4739 | norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e)" | 
| 53634 | 4740 | (is "?l = ?r") | 
| 4741 | proof | |
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4742 | assume rhs: ?r | 
| 53634 | 4743 | show ?l | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4744 | proof (subst has_integral', intro allI impI) | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4745 | fix e::real | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4746 | assume "e > 0" | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4747 | from rhs[THEN conjunct2,rule_format,OF this] | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4748 | show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4749 | (\<exists>z. ((\<lambda>x. if x \<in> s then f x else 0) has_integral z) | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4750 | (cbox a b) \<and> norm (z - i) < e)" | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4751 | apply (rule ex_forward) | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4752 | using rhs by blast | 
| 53634 | 4753 | qed | 
| 4754 | next | |
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4755 | let ?\<Phi> = "\<lambda>e a b. \<exists>z. ((\<lambda>x. if x \<in> s then f x else 0) has_integral z) (cbox a b) \<and> norm (z - i) < e" | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4756 | assume ?l | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4757 | then have lhs: "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> ?\<Phi> e a b" if "e > 0" for e | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4758 | using that has_integral'[of f] by auto | 
| 35751 | 4759 | let ?f = "\<lambda>x. if x \<in> s then f x else 0" | 
| 53634 | 4760 | show ?r | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4761 | proof (intro conjI allI impI) | 
| 53634 | 4762 | fix a b :: 'n | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4763 | from lhs[OF zero_less_one] | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4764 | obtain B where "0 < B" and B: "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> ?\<Phi> 1 a b" | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4765 | by blast | 
| 53634 | 4766 | let ?a = "\<Sum>i\<in>Basis. min (a\<bullet>i) (-B) *\<^sub>R i::'n" | 
| 4767 | let ?b = "\<Sum>i\<in>Basis. max (b\<bullet>i) B *\<^sub>R i::'n" | |
| 56188 | 4768 | show "?f integrable_on cbox a b" | 
| 53634 | 4769 | proof (rule integrable_subinterval[of _ ?a ?b]) | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4770 | have "?a \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> ?b \<bullet> i" if "norm (0 - x) < B" "i \<in> Basis" for x i | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4771 | using Basis_le_norm[of i x] that by (auto simp add:field_simps) | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4772 | then have "ball 0 B \<subseteq> cbox ?a ?b" | 
| 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4773 | by (auto simp: mem_box dist_norm) | 
| 56188 | 4774 | then show "?f integrable_on cbox ?a ?b" | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4775 | unfolding integrable_on_def using B by blast | 
| 56188 | 4776 | show "cbox a b \<subseteq> cbox ?a ?b" | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4777 | by (force simp: mem_box) | 
| 53634 | 4778 | qed | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4779 | |
| 53634 | 4780 | fix e :: real | 
| 4781 | assume "e > 0" | |
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4782 | with lhs show "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> | 
| 56188 | 4783 | norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - i) < e" | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 4784 | by (metis (no_types, lifting) has_integral_integrable_integral) | 
| 53634 | 4785 | qed | 
| 4786 | qed | |
| 35751 | 4787 | |
| 35752 | 4788 | |
| 60420 | 4789 | subsection \<open>Continuity of the integral (for a 1-dimensional interval).\<close> | 
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 4790 | |
| 53634 | 4791 | lemma integrable_alt: | 
| 56188 | 4792 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 53634 | 4793 | shows "f integrable_on s \<longleftrightarrow> | 
| 56188 | 4794 | (\<forall>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b) \<and> | 
| 4795 | (\<forall>e>0. \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow> | |
| 4796 | norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - | |
| 4797 | integral (cbox c d) (\<lambda>x. if x \<in> s then f x else 0)) < e)" | |
| 53634 | 4798 | (is "?l = ?r") | 
| 4799 | proof | |
| 66505 | 4800 | let ?F = "\<lambda>x. if x \<in> s then f x else 0" | 
| 53634 | 4801 | assume ?l | 
| 66505 | 4802 | then obtain y where intF: "\<And>a b. ?F integrable_on cbox a b" | 
| 4803 | and y: "\<And>e. 0 < e \<Longrightarrow> | |
| 4804 | \<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> norm (integral (cbox a b) ?F - y) < e" | |
| 4805 | unfolding integrable_on_def has_integral_alt'[of f] by auto | |
| 53634 | 4806 | show ?r | 
| 66505 | 4807 | proof (intro conjI allI impI intF) | 
| 4808 | fix e::real | |
| 4809 | assume "e > 0" | |
| 53634 | 4810 | then have "e/2 > 0" | 
| 4811 | by auto | |
| 66505 | 4812 | obtain B where "0 < B" | 
| 4813 | and B: "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> norm (integral (cbox a b) ?F - y) < e/2" | |
| 4814 | using \<open>0 < e/2\<close> y by blast | |
| 4815 | show "\<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow> | |
| 4816 | norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" | |
| 4817 | proof (intro conjI exI impI allI, rule \<open>0 < B\<close>) | |
| 4818 | fix a b c d::'n | |
| 4819 | assume sub: "ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d" | |
| 4820 | show "norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" | |
| 4821 | using sub by (auto intro: norm_triangle_half_l dest: B) | |
| 53634 | 4822 | qed | 
| 4823 | qed | |
| 4824 | next | |
| 66505 | 4825 | let ?F = "\<lambda>x. if x \<in> s then f x else 0" | 
| 4826 | assume rhs: ?r | |
| 56188 | 4827 | let ?cube = "\<lambda>n. cbox (\<Sum>i\<in>Basis. - real n *\<^sub>R i::'n) (\<Sum>i\<in>Basis. real n *\<^sub>R i)" | 
| 66508 | 4828 | have "Cauchy (\<lambda>n. integral (?cube n) ?F)" | 
| 4829 | unfolding Cauchy_def | |
| 4830 | proof (intro allI impI) | |
| 4831 | fix e::real | |
| 4832 | assume "e > 0" | |
| 66505 | 4833 | with rhs obtain B where "0 < B" | 
| 4834 | and B: "\<And>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d | |
| 4835 | \<Longrightarrow> norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e" | |
| 4836 | by blast | |
| 4837 | obtain N where N: "B \<le> real N" | |
| 4838 | using real_arch_simple by blast | |
| 66508 | 4839 | have "ball 0 B \<subseteq> ?cube n" if n: "n \<ge> N" for n | 
| 4840 | proof - | |
| 67399 | 4841 | have "sum (( *\<^sub>R) (- real n)) Basis \<bullet> i \<le> x \<bullet> i \<and> | 
| 4842 | x \<bullet> i \<le> sum (( *\<^sub>R) (real n)) Basis \<bullet> i" | |
| 66505 | 4843 | if "norm x < B" "i \<in> Basis" for x i::'n | 
| 4844 | using Basis_le_norm[of i x] n N that by (auto simp add: field_simps sum_negf) | |
| 66508 | 4845 | then show ?thesis | 
| 66505 | 4846 | by (auto simp: mem_box dist_norm) | 
| 66508 | 4847 | qed | 
| 4848 | then show "\<exists>M. \<forall>m\<ge>M. \<forall>n\<ge>M. dist (integral (?cube m) ?F) (integral (?cube n) ?F) < e" | |
| 66505 | 4849 | by (fastforce simp add: dist_norm intro!: B) | 
| 53634 | 4850 | qed | 
| 66508 | 4851 | then obtain i where i: "(\<lambda>n. integral (?cube n) ?F) \<longlonglongrightarrow> i" | 
| 4852 | using convergent_eq_Cauchy by blast | |
| 4853 | have "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> norm (integral (cbox a b) ?F - i) < e" | |
| 4854 | if "e > 0" for e | |
| 4855 | proof - | |
| 4856 | have *: "e/2 > 0" using that by auto | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 4857 | then obtain N where N: "\<And>n. N \<le> n \<Longrightarrow> norm (i - integral (?cube n) ?F) < e/2" | 
| 66508 | 4858 | using i[THEN LIMSEQ_D, simplified norm_minus_commute] by meson | 
| 4859 | obtain B where "0 < B" | |
| 4860 | and B: "\<And>a b c d. \<lbrakk>ball 0 B \<subseteq> cbox a b; ball 0 B \<subseteq> cbox c d\<rbrakk> \<Longrightarrow> | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 4861 | norm (integral (cbox a b) ?F - integral (cbox c d) ?F) < e/2" | 
| 66508 | 4862 | using rhs * by meson | 
| 53634 | 4863 | let ?B = "max (real N) B" | 
| 66508 | 4864 | show ?thesis | 
| 4865 | proof (intro exI conjI allI impI) | |
| 53634 | 4866 | show "0 < ?B" | 
| 66508 | 4867 | using \<open>B > 0\<close> by auto | 
| 53634 | 4868 | fix a b :: 'n | 
| 66508 | 4869 | assume "ball 0 ?B \<subseteq> cbox a b" | 
| 4870 | moreover obtain n where n: "max (real N) B \<le> real n" | |
| 4871 | using real_arch_simple by blast | |
| 4872 | moreover have "ball 0 B \<subseteq> ?cube n" | |
| 4873 | proof | |
| 53634 | 4874 | fix x :: 'n | 
| 4875 | assume x: "x \<in> ball 0 B" | |
| 66508 | 4876 | have "\<lbrakk>norm (0 - x) < B; i \<in> Basis\<rbrakk> | 
| 67399 | 4877 | \<Longrightarrow> sum (( *\<^sub>R) (-n)) Basis \<bullet> i\<le> x \<bullet> i \<and> x \<bullet> i \<le> sum (( *\<^sub>R) n) Basis \<bullet> i" for i | 
| 66508 | 4878 | using Basis_le_norm[of i x] n by (auto simp add: field_simps sum_negf) | 
| 4879 | then show "x \<in> ?cube n" | |
| 4880 | using x by (auto simp: mem_box dist_norm) | |
| 4881 | qed | |
| 4882 | ultimately show "norm (integral (cbox a b) ?F - i) < e" | |
| 4883 | using norm_triangle_half_l [OF B N] by force | |
| 53634 | 4884 | qed | 
| 4885 | qed | |
| 66508 | 4886 | then show ?l unfolding integrable_on_def has_integral_alt'[of f] | 
| 4887 | using rhs by blast | |
| 53634 | 4888 | qed | 
| 4889 | ||
| 4890 | lemma integrable_altD: | |
| 56188 | 4891 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 4892 | assumes "f integrable_on s" | 
| 56188 | 4893 | shows "\<And>a b. (\<lambda>x. if x \<in> s then f x else 0) integrable_on cbox a b" | 
| 4894 | and "\<And>e. e > 0 \<Longrightarrow> \<exists>B>0. \<forall>a b c d. ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow> | |
| 4895 | norm (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - integral (cbox c d) (\<lambda>x. if x \<in> s then f x else 0)) < e" | |
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 4896 | using assms[unfolded integrable_alt[of f]] by auto | 
| 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 4897 | |
| 56188 | 4898 | lemma integrable_on_subcbox: | 
| 4899 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | |
| 66505 | 4900 | assumes intf: "f integrable_on S" | 
| 4901 | and sub: "cbox a b \<subseteq> S" | |
| 56188 | 4902 | shows "f integrable_on cbox a b" | 
| 66505 | 4903 | proof - | 
| 4904 | have "(\<lambda>x. if x \<in> S then f x else 0) integrable_on cbox a b" | |
| 4905 | by (simp add: intf integrable_altD(1)) | |
| 4906 | then show ?thesis | |
| 4907 | by (metis (mono_tags) sub integrable_restrict_Int le_inf_iff order_refl subset_antisym) | |
| 4908 | qed | |
| 53634 | 4909 | |
| 4910 | ||
| 60420 | 4911 | subsection \<open>A straddling criterion for integrability\<close> | 
| 53634 | 4912 | |
| 4913 | lemma integrable_straddle_interval: | |
| 56188 | 4914 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4915 | assumes "\<And>e. e>0 \<Longrightarrow> \<exists>g h i j. (g has_integral i) (cbox a b) \<and> (h has_integral j) (cbox a b) \<and> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4916 | \<bar>i - j\<bar> < e \<and> (\<forall>x\<in>cbox a b. (g x) \<le> f x \<and> f x \<le> h x)" | 
| 56188 | 4917 | shows "f integrable_on cbox a b" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4918 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4919 | have "\<exists>d. gauge d \<and> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4920 | (\<forall>p1 p2. p1 tagged_division_of cbox a b \<and> d fine p1 \<and> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4921 | p2 tagged_division_of cbox a b \<and> d fine p2 \<longrightarrow> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4922 | \<bar>(\<Sum>(x,K)\<in>p1. content K *\<^sub>R f x) - (\<Sum>(x,K)\<in>p2. content K *\<^sub>R f x)\<bar> < e)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4923 | if "e > 0" for e | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4924 | proof - | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4925 | have e: "e/3 > 0" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4926 | using that by auto | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4927 | then obtain g h i j where ij: "\<bar>i - j\<bar> < e/3" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4928 | and "(g has_integral i) (cbox a b)" | 
| 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4929 | and "(h has_integral j) (cbox a b)" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4930 | and fgh: "\<And>x. x \<in> cbox a b \<Longrightarrow> g x \<le> f x \<and> f x \<le> h x" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4931 | using assms real_norm_def by metis | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4932 | then obtain d1 d2 where "gauge d1" "gauge d2" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4933 | and d1: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; d1 fine p\<rbrakk> \<Longrightarrow> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4934 | \<bar>(\<Sum>(x,K)\<in>p. content K *\<^sub>R g x) - i\<bar> < e/3" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4935 | and d2: "\<And>p. \<lbrakk>p tagged_division_of cbox a b; d2 fine p\<rbrakk> \<Longrightarrow> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4936 | \<bar>(\<Sum>(x,K) \<in> p. content K *\<^sub>R h x) - j\<bar> < e/3" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4937 | by (metis e has_integral real_norm_def) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4938 | have "\<bar>(\<Sum>(x,K) \<in> p1. content K *\<^sub>R f x) - (\<Sum>(x,K) \<in> p2. content K *\<^sub>R f x)\<bar> < e" | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4939 | if p1: "p1 tagged_division_of cbox a b" and 11: "d1 fine p1" and 21: "d2 fine p1" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4940 | and p2: "p2 tagged_division_of cbox a b" and 12: "d1 fine p2" and 22: "d2 fine p2" for p1 p2 | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4941 | proof - | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 4942 | have *: "\<And>g1 g2 h1 h2 f1 f2. | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4943 | \<lbrakk>\<bar>g2 - i\<bar> < e/3; \<bar>g1 - i\<bar> < e/3; \<bar>h2 - j\<bar> < e/3; \<bar>h1 - j\<bar> < e/3; | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4944 | g1 - h2 \<le> f1 - f2; f1 - f2 \<le> h1 - g2\<rbrakk> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4945 | \<Longrightarrow> \<bar>f1 - f2\<bar> < e" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4946 | using \<open>e > 0\<close> ij by arith | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4947 | have 0: "(\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p1. content k *\<^sub>R g x) \<ge> 0" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4948 | "0 \<le> (\<Sum>(x, k)\<in>p2. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4949 | "(\<Sum>(x, k)\<in>p2. content k *\<^sub>R f x) - (\<Sum>(x, k)\<in>p2. content k *\<^sub>R g x) \<ge> 0" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4950 | "0 \<le> (\<Sum>(x, k)\<in>p1. content k *\<^sub>R h x) - (\<Sum>(x, k)\<in>p1. content k *\<^sub>R f x)" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4951 | unfolding sum_subtractf[symmetric] | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4952 | apply (auto intro!: sum_nonneg) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4953 | apply (meson fgh measure_nonneg mult_left_mono tag_in_interval that sum_nonneg)+ | 
| 53634 | 4954 | done | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4955 | show ?thesis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4956 | proof (rule *) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4957 | show "\<bar>(\<Sum>(x,K) \<in> p2. content K *\<^sub>R g x) - i\<bar> < e/3" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4958 | by (rule d1[OF p2 12]) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4959 | show "\<bar>(\<Sum>(x,K) \<in> p1. content K *\<^sub>R g x) - i\<bar> < e/3" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4960 | by (rule d1[OF p1 11]) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4961 | show "\<bar>(\<Sum>(x,K) \<in> p2. content K *\<^sub>R h x) - j\<bar> < e/3" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4962 | by (rule d2[OF p2 22]) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4963 | show "\<bar>(\<Sum>(x,K) \<in> p1. content K *\<^sub>R h x) - j\<bar> < e/3" | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4964 | by (rule d2[OF p1 21]) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4965 | qed (use 0 in auto) | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4966 | qed | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4967 | then show ?thesis | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4968 | by (rule_tac x="\<lambda>x. d1 x \<inter> d2 x" in exI) | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4969 | (auto simp: fine_Int intro: \<open>gauge d1\<close> \<open>gauge d2\<close> d1 d2) | 
| 53634 | 4970 | qed | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4971 | then show ?thesis | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4972 | by (simp add: integrable_Cauchy) | 
| 53634 | 4973 | qed | 
| 53399 | 4974 | |
| 53638 | 4975 | lemma integrable_straddle: | 
| 56188 | 4976 | fixes f :: "'n::euclidean_space \<Rightarrow> real" | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4977 | assumes "\<And>e. e>0 \<Longrightarrow> \<exists>g h i j. (g has_integral i) s \<and> (h has_integral j) s \<and> | 
| 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 4978 | \<bar>i - j\<bar> < e \<and> (\<forall>x\<in>s. g x \<le> f x \<and> f x \<le> h x)" | 
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 4979 | shows "f integrable_on s" | 
| 53638 | 4980 | proof - | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4981 | let ?fs = "(\<lambda>x. if x \<in> s then f x else 0)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4982 | have "?fs integrable_on cbox a b" for a b | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4983 | proof (rule integrable_straddle_interval) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4984 | fix e::real | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4985 | assume "e > 0" | 
| 53638 | 4986 | then have *: "e/4 > 0" | 
| 4987 | by auto | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4988 | with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4989 | and ij: "\<bar>i - j\<bar> < e/4" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4990 | and fgh: "\<And>x. x \<in> s \<Longrightarrow> g x \<le> f x \<and> f x \<le> h x" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4991 | by metis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4992 | let ?gs = "(\<lambda>x. if x \<in> s then g x else 0)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4993 | let ?hs = "(\<lambda>x. if x \<in> s then h x else 0)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4994 | obtain Bg where Bg: "\<And>a b. ball 0 Bg \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox a b) ?gs - i\<bar> < e/4" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4995 | and int_g: "\<And>a b. ?gs integrable_on cbox a b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4996 | using g * unfolding has_integral_alt' real_norm_def by meson | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4997 | obtain Bh where | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4998 | Bh: "\<And>a b. ball 0 Bh \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox a b) ?hs - j\<bar> < e/4" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 4999 | and int_h: "\<And>a b. ?hs integrable_on cbox a b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5000 | using h * unfolding has_integral_alt' real_norm_def by meson | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5001 | define c where "c = (\<Sum>i\<in>Basis. min (a\<bullet>i) (- (max Bg Bh)) *\<^sub>R i)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5002 | define d where "d = (\<Sum>i\<in>Basis. max (b\<bullet>i) (max Bg Bh) *\<^sub>R i)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5003 | have "\<lbrakk>norm (0 - x) < Bg; i \<in> Basis\<rbrakk> \<Longrightarrow> c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" for x i | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5004 | using Basis_le_norm[of i x] unfolding c_def d_def by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5005 | then have ballBg: "ball 0 Bg \<subseteq> cbox c d" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5006 | by (auto simp: mem_box dist_norm) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5007 | have "\<lbrakk>norm (0 - x) < Bh; i \<in> Basis\<rbrakk> \<Longrightarrow> c \<bullet> i \<le> x \<bullet> i \<and> x \<bullet> i \<le> d \<bullet> i" for x i | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5008 | using Basis_le_norm[of i x] unfolding c_def d_def by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5009 | then have ballBh: "ball 0 Bh \<subseteq> cbox c d" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5010 | by (auto simp: mem_box dist_norm) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5011 | have ab_cd: "cbox a b \<subseteq> cbox c d" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5012 | by (auto simp: c_def d_def subset_box_imp) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5013 | have **: "\<And>ch cg ag ah::real. \<lbrakk>\<bar>ah - ag\<bar> \<le> \<bar>ch - cg\<bar>; \<bar>cg - i\<bar> < e/4; \<bar>ch - j\<bar> < e/4\<rbrakk> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5014 | \<Longrightarrow> \<bar>ag - ah\<bar> < e" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5015 | using ij by arith | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5016 | show "\<exists>g h i j. (g has_integral i) (cbox a b) \<and> (h has_integral j) (cbox a b) \<and> \<bar>i - j\<bar> < e \<and> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5017 | (\<forall>x\<in>cbox a b. g x \<le> (if x \<in> s then f x else 0) \<and> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5018 | (if x \<in> s then f x else 0) \<le> h x)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5019 | proof (intro exI ballI conjI) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5020 | have eq: "\<And>x f g. (if x \<in> s then f x else 0) - (if x \<in> s then g x else 0) = | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5021 | (if x \<in> s then f x - g x else (0::real))" | 
| 53638 | 5022 | by auto | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5023 | have int_hg: "(\<lambda>x. if x \<in> s then h x - g x else 0) integrable_on cbox a b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5024 | "(\<lambda>x. if x \<in> s then h x - g x else 0) integrable_on cbox c d" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5025 | by (metis (no_types) integrable_diff g h has_integral_integrable integrable_altD(1))+ | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5026 | show "(?gs has_integral integral (cbox a b) ?gs) (cbox a b)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5027 | "(?hs has_integral integral (cbox a b) ?hs) (cbox a b)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5028 | by (intro integrable_integral int_g int_h)+ | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5029 | then have "integral (cbox a b) ?gs \<le> integral (cbox a b) ?hs" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5030 | apply (rule has_integral_le) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5031 | using fgh by force | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5032 | then have "0 \<le> integral (cbox a b) ?hs - integral (cbox a b) ?gs" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5033 | by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5034 | then have "\<bar>integral (cbox a b) ?hs - integral (cbox a b) ?gs\<bar> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5035 | \<le> \<bar>integral (cbox c d) ?hs - integral (cbox c d) ?gs\<bar>" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5036 | apply (simp add: integral_diff [symmetric] int_g int_h) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5037 | apply (subst abs_of_nonneg[OF integral_nonneg[OF integrable_diff, OF int_h int_g]]) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5038 | using fgh apply (force simp: eq intro!: integral_subset_le [OF ab_cd int_hg])+ | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5039 | done | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5040 | then show "\<bar>integral (cbox a b) ?gs - integral (cbox a b) ?hs\<bar> < e" | 
| 53638 | 5041 | apply (rule **) | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5042 | apply (rule Bg ballBg Bh ballBh)+ | 
| 53638 | 5043 | done | 
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5044 | show "\<And>x. x \<in> cbox a b \<Longrightarrow> ?gs x \<le> ?fs x" "\<And>x. x \<in> cbox a b \<Longrightarrow> ?fs x \<le> ?hs x" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5045 | using fgh by auto | 
| 53638 | 5046 | qed | 
| 5047 | qed | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5048 | then have int_f: "?fs integrable_on cbox a b" for a b | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5049 | by simp | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5050 | have "\<exists>B>0. \<forall>a b c d. | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5051 | ball 0 B \<subseteq> cbox a b \<and> ball 0 B \<subseteq> cbox c d \<longrightarrow> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5052 | abs (integral (cbox a b) ?fs - integral (cbox c d) ?fs) < e" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5053 | if "0 < e" for e | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5054 | proof - | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5055 | have *: "e/3 > 0" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5056 | using that by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5057 | with assms obtain g h i j where g: "(g has_integral i) s" and h: "(h has_integral j) s" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5058 | and ij: "\<bar>i - j\<bar> < e/3" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5059 | and fgh: "\<And>x. x \<in> s \<Longrightarrow> g x \<le> f x \<and> f x \<le> h x" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5060 | by metis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5061 | let ?gs = "(\<lambda>x. if x \<in> s then g x else 0)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5062 | let ?hs = "(\<lambda>x. if x \<in> s then h x else 0)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5063 | obtain Bg where "Bg > 0" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5064 | and Bg: "\<And>a b. ball 0 Bg \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox a b) ?gs - i\<bar> < e/3" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5065 | and int_g: "\<And>a b. ?gs integrable_on cbox a b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5066 | using g * unfolding has_integral_alt' real_norm_def by meson | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5067 | obtain Bh where "Bh > 0" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5068 | and Bh: "\<And>a b. ball 0 Bh \<subseteq> cbox a b \<Longrightarrow> \<bar>integral (cbox a b) ?hs - j\<bar> < e/3" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5069 | and int_h: "\<And>a b. ?hs integrable_on cbox a b" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5070 | using h * unfolding has_integral_alt' real_norm_def by meson | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5071 |     { fix a b c d :: 'n
 | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5072 | assume as: "ball 0 (max Bg Bh) \<subseteq> cbox a b" "ball 0 (max Bg Bh) \<subseteq> cbox c d" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5073 | have **: "ball 0 Bg \<subseteq> ball (0::'n) (max Bg Bh)" "ball 0 Bh \<subseteq> ball (0::'n) (max Bg Bh)" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5074 | by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5075 | have *: "\<And>ga gc ha hc fa fc. \<lbrakk>\<bar>ga - i\<bar> < e/3; \<bar>gc - i\<bar> < e/3; \<bar>ha - j\<bar> < e/3; | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5076 | \<bar>hc - j\<bar> < e/3; ga \<le> fa; fa \<le> ha; gc \<le> fc; fc \<le> hc\<rbrakk> \<Longrightarrow> | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5077 | \<bar>fa - fc\<bar> < e" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5078 | using ij by arith | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5079 | have "abs (integral (cbox a b) (\<lambda>x. if x \<in> s then f x else 0) - integral (cbox c d) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5080 | (\<lambda>x. if x \<in> s then f x else 0)) < e" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5081 | proof (rule *) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5082 | show "\<bar>integral (cbox a b) ?gs - i\<bar> < e/3" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5083 | using "**" Bg as by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5084 | show "\<bar>integral (cbox c d) ?gs - i\<bar> < e/3" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5085 | using "**" Bg as by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5086 | show "\<bar>integral (cbox a b) ?hs - j\<bar> < e/3" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5087 | using "**" Bh as by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5088 | show "\<bar>integral (cbox c d) ?hs - j\<bar> < e/3" | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5089 | using "**" Bh as by blast | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5090 | qed (use int_f int_g int_h fgh in \<open>simp_all add: integral_le\<close>) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5091 | } | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5092 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5093 | apply (rule_tac x="max Bg Bh" in exI) | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5094 | using \<open>Bg > 0\<close> by auto | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5095 | qed | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5096 | then show ?thesis | 
| 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 5097 | unfolding integrable_alt[of f] real_norm_def by (blast intro: int_f) | 
| 53638 | 5098 | qed | 
| 5099 | ||
| 5100 | ||
| 60420 | 5101 | subsection \<open>Adding integrals over several sets\<close> | 
| 53638 | 5102 | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5103 | lemma has_integral_Un: | 
| 56188 | 5104 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5105 | assumes f: "(f has_integral i) s" "(f has_integral j) t" | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5106 | and neg: "negligible (s \<inter> t)" | 
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 5107 | shows "(f has_integral (i + j)) (s \<union> t)" | 
| 53638 | 5108 | proof - | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 5109 | note * = has_integral_restrict_UNIV[symmetric, of f] | 
| 53638 | 5110 | show ?thesis | 
| 5111 | unfolding * | |
| 5112 | apply (rule has_integral_spike[OF assms(3)]) | |
| 5113 | defer | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5114 | apply (rule has_integral_add[OF f[unfolded *]]) | 
| 53638 | 5115 | apply auto | 
| 5116 | done | |
| 5117 | qed | |
| 5118 | ||
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5119 | lemma integrable_Un: | 
| 63296 | 5120 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: banach" | 
| 5121 | assumes "negligible (A \<inter> B)" "f integrable_on A" "f integrable_on B" | |
| 5122 | shows "f integrable_on (A \<union> B)" | |
| 5123 | proof - | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 5124 | from assms obtain y z where "(f has_integral y) A" "(f has_integral z) B" | 
| 63296 | 5125 | by (auto simp: integrable_on_def) | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5126 | from has_integral_Un[OF this assms(1)] show ?thesis by (auto simp: integrable_on_def) | 
| 63296 | 5127 | qed | 
| 5128 | ||
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5129 | lemma integrable_Un': | 
| 63296 | 5130 | fixes f :: "'a::euclidean_space \<Rightarrow> 'b :: banach" | 
| 5131 | assumes "f integrable_on A" "f integrable_on B" "negligible (A \<inter> B)" "C = A \<union> B" | |
| 5132 | shows "f integrable_on C" | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5133 | using integrable_Un[of A B f] assms by simp | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5134 | |
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5135 | lemma has_integral_Union: | 
| 56188 | 5136 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66560 | 5137 | assumes \<T>: "finite \<T>" | 
| 5138 | and int: "\<And>S. S \<in> \<T> \<Longrightarrow> (f has_integral (i S)) S" | |
| 5139 | and neg: "\<And>S S'. \<lbrakk>S \<in> \<T>; S' \<in> \<T>; S \<noteq> S'\<rbrakk> \<Longrightarrow> negligible (S \<inter> S')" | |
| 5140 | shows "(f has_integral (sum i \<T>)) (\<Union>\<T>)" | |
| 53638 | 5141 | proof - | 
| 66560 | 5142 |   let ?\<U> = "((\<lambda>(a,b). a \<inter> b) ` {(a,b). a \<in> \<T> \<and> b \<in> {y. y \<in> \<T> \<and> a \<noteq> y}})"
 | 
| 5143 | have "((\<lambda>x. if x \<in> \<Union>\<T> then f x else 0) has_integral sum i \<T>) UNIV" | |
| 5144 | proof (rule has_integral_spike) | |
| 5145 | show "negligible (\<Union>?\<U>)" | |
| 5146 | proof (rule negligible_Union) | |
| 5147 | have "finite (\<T> \<times> \<T>)" | |
| 5148 | by (simp add: \<T>) | |
| 5149 |       moreover have "{(a, b). a \<in> \<T> \<and> b \<in> {y \<in> \<T>. a \<noteq> y}} \<subseteq> \<T> \<times> \<T>"
 | |
| 5150 | by auto | |
| 5151 | ultimately show "finite ?\<U>" | |
| 5152 | by (blast intro: finite_subset[of _ "\<T> \<times> \<T>"]) | |
| 5153 | show "\<And>t. t \<in> ?\<U> \<Longrightarrow> negligible t" | |
| 5154 | using neg by auto | |
| 5155 | qed | |
| 5156 | next | |
| 5157 | show "(if x \<in> \<Union>\<T> then f x else 0) = (\<Sum>A\<in>\<T>. if x \<in> A then f x else 0)" | |
| 5158 | if "x \<in> UNIV - (\<Union>?\<U>)" for x | |
| 5159 | proof clarsimp | |
| 5160 | fix S assume "S \<in> \<T>" "x \<in> S" | |
| 5161 | moreover then have "\<forall>b\<in>\<T>. x \<in> b \<longleftrightarrow> b = S" | |
| 5162 | using that by blast | |
| 5163 | ultimately show "f x = (\<Sum>A\<in>\<T>. if x \<in> A then f x else 0)" | |
| 5164 | by (simp add: sum.delta[OF \<T>]) | |
| 5165 | qed | |
| 5166 | next | |
| 5167 | show "((\<lambda>x. \<Sum>A\<in>\<T>. if x \<in> A then f x else 0) has_integral (\<Sum>A\<in>\<T>. i A)) UNIV" | |
| 5168 | apply (rule has_integral_sum [OF \<T>]) | |
| 5169 | using int by (simp add: has_integral_restrict_UNIV) | |
| 5170 | qed | |
| 61165 | 5171 | then show ?thesis | 
| 66560 | 5172 | using has_integral_restrict_UNIV by blast | 
| 53638 | 5173 | qed | 
| 5174 | ||
| 5175 | ||
| 60420 | 5176 | text \<open>In particular adding integrals over a division, maybe not of an interval.\<close> | 
| 53638 | 5177 | |
| 5178 | lemma has_integral_combine_division: | |
| 56188 | 5179 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66561 | 5180 | assumes "\<D> division_of S" | 
| 5181 | and "\<And>k. k \<in> \<D> \<Longrightarrow> (f has_integral (i k)) k" | |
| 5182 | shows "(f has_integral (sum i \<D>)) S" | |
| 53638 | 5183 | proof - | 
| 66561 | 5184 | note \<D> = division_ofD[OF assms(1)] | 
| 5185 | have neg: "negligible (S \<inter> s')" if "S \<in> \<D>" "s' \<in> \<D>" "S \<noteq> s'" for S s' | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5186 | proof - | 
| 66561 | 5187 | obtain a c b \<D> where obt: "S = cbox a b" "s' = cbox c \<D>" | 
| 5188 | by (meson \<open>S \<in> \<D>\<close> \<open>s' \<in> \<D>\<close> \<D>(4)) | |
| 5189 | from \<D>(5)[OF that] show ?thesis | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5190 | unfolding obt interior_cbox | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5191 | by (metis (no_types, lifting) Diff_empty Int_interval box_Int_box negligible_frontier_interval) | 
| 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5192 | qed | 
| 53638 | 5193 | show ?thesis | 
| 66561 | 5194 | unfolding \<D>(6)[symmetric] | 
| 5195 | by (auto intro: \<D> neg assms has_integral_Union) | |
| 53638 | 5196 | qed | 
| 5197 | ||
| 5198 | lemma integral_combine_division_bottomup: | |
| 56188 | 5199 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66561 | 5200 | assumes "\<D> division_of S" "\<And>k. k \<in> \<D> \<Longrightarrow> f integrable_on k" | 
| 5201 | shows "integral S f = sum (\<lambda>i. integral i f) \<D>" | |
| 53638 | 5202 | apply (rule integral_unique) | 
| 66561 | 5203 | by (meson assms has_integral_combine_division has_integral_integrable_integral) | 
| 53638 | 5204 | |
| 5205 | lemma has_integral_combine_division_topdown: | |
| 56188 | 5206 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66561 | 5207 | assumes f: "f integrable_on S" | 
| 5208 | and \<D>: "\<D> division_of K" | |
| 5209 | and "K \<subseteq> S" | |
| 5210 | shows "(f has_integral (sum (\<lambda>i. integral i f) \<D>)) K" | |
| 5211 | proof - | |
| 5212 | have "f integrable_on L" if "L \<in> \<D>" for L | |
| 5213 | proof - | |
| 5214 | have "L \<subseteq> S" | |
| 5215 | using \<open>K \<subseteq> S\<close> \<D> that by blast | |
| 5216 | then show "f integrable_on L" | |
| 5217 | using that by (metis (no_types) f \<D> division_ofD(4) integrable_on_subcbox) | |
| 5218 | qed | |
| 5219 | then show ?thesis | |
| 5220 | by (meson \<D> has_integral_combine_division has_integral_integrable_integral) | |
| 53638 | 5221 | qed | 
| 5222 | ||
| 5223 | lemma integral_combine_division_topdown: | |
| 56188 | 5224 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5225 | assumes "f integrable_on S" | 
| 66561 | 5226 | and "\<D> division_of S" | 
| 5227 | shows "integral S f = sum (\<lambda>i. integral i f) \<D>" | |
| 53638 | 5228 | apply (rule integral_unique) | 
| 5229 | apply (rule has_integral_combine_division_topdown) | |
| 5230 | using assms | |
| 5231 | apply auto | |
| 5232 | done | |
| 5233 | ||
| 5234 | lemma integrable_combine_division: | |
| 56188 | 5235 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66561 | 5236 | assumes \<D>: "\<D> division_of S" | 
| 5237 | and f: "\<And>i. i \<in> \<D> \<Longrightarrow> f integrable_on i" | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5238 | shows "f integrable_on S" | 
| 66561 | 5239 | using f unfolding integrable_on_def by (metis has_integral_combine_division[OF \<D>]) | 
| 53638 | 5240 | |
| 5241 | lemma integrable_on_subdivision: | |
| 56188 | 5242 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66561 | 5243 | assumes \<D>: "\<D> division_of i" | 
| 5244 | and f: "f integrable_on S" | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5245 | and "i \<subseteq> S" | 
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 5246 | shows "f integrable_on i" | 
| 66561 | 5247 | proof - | 
| 5248 | have "f integrable_on i" if "i \<in> \<D>" for i | |
| 5249 | proof - | |
| 5250 | have "i \<subseteq> S" | |
| 5251 | using assms that by auto | |
| 5252 | then show "f integrable_on i" | |
| 5253 | using that by (metis (no_types) \<D> f division_ofD(4) integrable_on_subcbox) | |
| 5254 | qed | |
| 5255 | then show ?thesis | |
| 5256 | using \<D> integrable_combine_division by blast | |
| 53638 | 5257 | qed | 
| 5258 | ||
| 5259 | ||
| 60420 | 5260 | subsection \<open>Also tagged divisions\<close> | 
| 53638 | 5261 | |
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5262 | lemma has_integral_iff: "(f has_integral i) S \<longleftrightarrow> (f integrable_on S \<and> integral S f = i)" | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 5263 | by blast | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 5264 | |
| 53638 | 5265 | lemma has_integral_combine_tagged_division: | 
| 56188 | 5266 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5267 | assumes "p tagged_division_of S" | 
| 53638 | 5268 | and "\<forall>(x,k) \<in> p. (f has_integral (i k)) k" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5269 | shows "(f has_integral (\<Sum>(x,k)\<in>p. i k)) S" | 
| 53638 | 5270 | proof - | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5271 | have *: "(f has_integral (\<Sum>k\<in>snd`p. integral k f)) S" | 
| 53638 | 5272 | using assms(2) | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 5273 | apply (intro has_integral_combine_division) | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 5274 | apply (auto simp: has_integral_integral[symmetric] intro: division_of_tagged_division[OF assms(1)]) | 
| 53638 | 5275 | apply auto | 
| 5276 | done | |
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 5277 | also have "(\<Sum>k\<in>snd`p. integral k f) = (\<Sum>(x, k)\<in>p. integral k f)" | 
| 64267 | 5278 | by (intro sum.over_tagged_division_lemma[OF assms(1), symmetric] integral_null) | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 5279 | (simp add: content_eq_0_interior) | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 5280 | finally show ?thesis | 
| 64267 | 5281 | using assms by (auto simp add: has_integral_iff intro!: sum.cong) | 
| 53638 | 5282 | qed | 
| 5283 | ||
| 5284 | lemma integral_combine_tagged_division_bottomup: | |
| 56188 | 5285 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 5286 | assumes "p tagged_division_of (cbox a b)" | |
| 53638 | 5287 | and "\<forall>(x,k)\<in>p. f integrable_on k" | 
| 64267 | 5288 | shows "integral (cbox a b) f = sum (\<lambda>(x,k). integral k f) p" | 
| 53638 | 5289 | apply (rule integral_unique) | 
| 5290 | apply (rule has_integral_combine_tagged_division[OF assms(1)]) | |
| 5291 | using assms(2) | |
| 5292 | apply auto | |
| 5293 | done | |
| 5294 | ||
| 5295 | lemma has_integral_combine_tagged_division_topdown: | |
| 56188 | 5296 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66561 | 5297 | assumes f: "f integrable_on cbox a b" | 
| 5298 | and p: "p tagged_division_of (cbox a b)" | |
| 5299 | shows "(f has_integral (sum (\<lambda>(x,K). integral K f) p)) (cbox a b)" | |
| 5300 | proof - | |
| 5301 | have "(f has_integral integral K f) K" if "(x,K) \<in> p" for x K | |
| 5302 | by (metis assms integrable_integral integrable_on_subcbox tagged_division_ofD(3,4) that) | |
| 5303 | then show ?thesis | |
| 5304 | by (metis assms case_prodI2 has_integral_integrable_integral integral_combine_tagged_division_bottomup) | |
| 53638 | 5305 | qed | 
| 5306 | ||
| 5307 | lemma integral_combine_tagged_division_topdown: | |
| 56188 | 5308 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 5309 | assumes "f integrable_on cbox a b" | |
| 5310 | and "p tagged_division_of (cbox a b)" | |
| 64267 | 5311 | shows "integral (cbox a b) f = sum (\<lambda>(x,k). integral k f) p" | 
| 66561 | 5312 | apply (rule integral_unique [OF has_integral_combine_tagged_division_topdown]) | 
| 5313 | using assms apply auto | |
| 53638 | 5314 | done | 
| 5315 | ||
| 5316 | ||
| 60420 | 5317 | subsection \<open>Henstock's lemma\<close> | 
| 53638 | 5318 | |
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5319 | lemma Henstock_lemma_part1: | 
| 56188 | 5320 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5321 | assumes intf: "f integrable_on cbox a b" | 
| 53638 | 5322 | and "e > 0" | 
| 5323 | and "gauge d" | |
| 66437 | 5324 | and less_e: "\<And>p. \<lbrakk>p tagged_division_of (cbox a b); d fine p\<rbrakk> \<Longrightarrow> | 
| 5325 | norm (sum (\<lambda>(x,K). content K *\<^sub>R f x) p - integral(cbox a b) f) < e" | |
| 56188 | 5326 | and p: "p tagged_partial_division_of (cbox a b)" "d fine p" | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5327 | shows "norm (sum (\<lambda>(x,K). content K *\<^sub>R f x - integral K f) p) \<le> e" (is "?lhs \<le> e") | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 5328 | proof (rule field_le_epsilon) | 
| 53638 | 5329 | fix k :: real | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5330 | assume "k > 0" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5331 | let ?SUM = "\<lambda>p. (\<Sum>(x,K) \<in> p. content K *\<^sub>R f x)" | 
| 53638 | 5332 | note p' = tagged_partial_division_ofD[OF p(1)] | 
| 56188 | 5333 | have "\<Union>(snd ` p) \<subseteq> cbox a b" | 
| 53638 | 5334 | using p'(3) by fastforce | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5335 | then obtain q where q: "snd ` p \<subseteq> q" and qdiv: "q division_of cbox a b" | 
| 66512 
89b6455b63b6
starting to unscramble bounded_variation_absolutely_integrable_interval
 paulson <lp15@cam.ac.uk> parents: 
66508diff
changeset | 5336 | by (meson p(1) partial_division_extend_interval partial_division_of_tagged_division) | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5337 | note q' = division_ofD[OF qdiv] | 
| 63040 | 5338 | define r where "r = q - snd ` p" | 
| 53638 | 5339 |   have "snd ` p \<inter> r = {}"
 | 
| 5340 | unfolding r_def by auto | |
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5341 | have "finite r" | 
| 53638 | 5342 | using q' unfolding r_def by auto | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5343 | have "\<exists>p. p tagged_division_of i \<and> d fine p \<and> | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5344 | norm (?SUM p - integral i f) < k / (real (card r) + 1)" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5345 | if "i\<in>r" for i | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5346 | proof - | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5347 | have gt0: "k / (real (card r) + 1) > 0" using \<open>k > 0\<close> by simp | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5348 | have i: "i \<in> q" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5349 | using that unfolding r_def by auto | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5350 | then obtain u v where uv: "i = cbox u v" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5351 | using q'(4) by blast | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5352 | then have "cbox u v \<subseteq> cbox a b" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5353 | using i q'(2) by auto | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5354 | then have "f integrable_on cbox u v" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5355 | by (rule integrable_subinterval[OF intf]) | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5356 | with integrable_integral[OF this, unfolded has_integral[of f]] | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5357 | obtain dd where "gauge dd" and dd: | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5358 | "\<And>\<D>. \<lbrakk>\<D> tagged_division_of cbox u v; dd fine \<D>\<rbrakk> \<Longrightarrow> | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5359 | norm (?SUM \<D> - integral (cbox u v) f) < k / (real (card r) + 1)" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5360 | using gt0 by auto | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5361 | with gauge_Int[OF \<open>gauge d\<close> \<open>gauge dd\<close>] | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5362 | obtain qq where qq: "qq tagged_division_of cbox u v" "(\<lambda>x. d x \<inter> dd x) fine qq" | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5363 | using fine_division_exists by blast | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5364 | with dd[of qq] show ?thesis | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5365 | by (auto simp: fine_Int uv) | 
| 53638 | 5366 | qed | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5367 | then obtain qq where qq: "\<And>i. i \<in> r \<Longrightarrow> qq i tagged_division_of i \<and> | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5368 | d fine qq i \<and> norm (?SUM (qq i) - integral i f) < k / (real (card r) + 1)" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5369 | by metis | 
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 5370 | |
| 53638 | 5371 | let ?p = "p \<union> \<Union>(qq ` r)" | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5372 | have "norm (?SUM ?p - integral (cbox a b) f) < e" | 
| 66437 | 5373 | proof (rule less_e) | 
| 53638 | 5374 | show "d fine ?p" | 
| 66437 | 5375 | by (metis (mono_tags, hide_lams) qq fine_Un fine_Union imageE p(2)) | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5376 | note ptag = tagged_partial_division_of_Union_self[OF p(1)] | 
| 52141 
eff000cab70f
weaker precendence of syntax for big intersection and union on sets
 haftmann parents: 
51642diff
changeset | 5377 | have "p \<union> \<Union>(qq ` r) tagged_division_of \<Union>(snd ` p) \<union> \<Union>r" | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5378 | proof (rule tagged_division_Un[OF ptag tagged_division_Union [OF \<open>finite r\<close>]]) | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5379 | show "\<And>i. i \<in> r \<Longrightarrow> qq i tagged_division_of i" | 
| 53638 | 5380 | using qq by auto | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5381 |       show "\<And>i1 i2. \<lbrakk>i1 \<in> r; i2 \<in> r; i1 \<noteq> i2\<rbrakk> \<Longrightarrow> interior i1 \<inter> interior i2 = {}"
 | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5382 | by (simp add: q'(5) r_def) | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5383 |       show "interior (UNION p snd) \<inter> interior (\<Union>r) = {}"
 | 
| 66299 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5384 | proof (rule Int_interior_Union_intervals [OF \<open>finite r\<close>]) | 
| 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5385 | show "open (interior (UNION p snd))" | 
| 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5386 | by blast | 
| 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5387 | show "\<And>T. T \<in> r \<Longrightarrow> \<exists>a b. T = cbox a b" | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5388 | by (simp add: q'(4) r_def) | 
| 66299 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5389 | have "finite (snd ` p)" | 
| 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5390 | by (simp add: p'(1)) | 
| 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5391 |         then show "\<And>T. T \<in> r \<Longrightarrow> interior (UNION p snd) \<inter> interior T = {}"
 | 
| 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5392 | apply (subst Int_commute) | 
| 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5393 | apply (rule Int_interior_Union_intervals) | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5394 | using r_def q'(5) q(1) apply auto | 
| 66299 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5395 | by (simp add: p'(4)) | 
| 
1b4aa3e3e4e6
partial cleanup of the horrible Tagged_Division
 paulson <lp15@cam.ac.uk> parents: 
66296diff
changeset | 5396 | qed | 
| 53638 | 5397 | qed | 
| 56188 | 5398 |     moreover have "\<Union>(snd ` p) \<union> \<Union>r = cbox a b" and "{qq i |i. i \<in> r} = qq ` r"
 | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5399 | using qdiv q unfolding Union_Un_distrib[symmetric] r_def by auto | 
| 56188 | 5400 | ultimately show "?p tagged_division_of (cbox a b)" | 
| 53638 | 5401 | by fastforce | 
| 5402 | qed | |
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5403 | then have "norm (?SUM p + (?SUM (\<Union>(qq ` r))) - integral (cbox a b) f) < e" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5404 | proof (subst sum.union_inter_neutral[symmetric, OF \<open>finite p\<close>], safe) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5405 | show "content L *\<^sub>R f x = 0" if "(x, L) \<in> p" "(x, L) \<in> qq K" "K \<in> r" for x K L | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5406 | proof - | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5407 | obtain u v where uv: "L = cbox u v" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5408 | using \<open>(x,L) \<in> p\<close> p'(4) by blast | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5409 | have "L \<subseteq> K" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5410 | using qq[OF that(3)] tagged_division_ofD(3) \<open>(x,L) \<in> qq K\<close> by metis | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5411 | have "L \<in> snd ` p" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5412 | using \<open>(x,L) \<in> p\<close> image_iff by fastforce | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5413 | then have "L \<in> q" "K \<in> q" "L \<noteq> K" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5414 | using that(1,3) q(1) unfolding r_def by auto | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5415 |       with q'(5) have "interior L = {}"
 | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5416 | using interior_mono[OF \<open>L \<subseteq> K\<close>] by blast | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5417 | then show "content L *\<^sub>R f x = 0" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5418 | unfolding uv content_eq_0_interior[symmetric] by auto | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5419 | qed | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5420 | show "finite (UNION r qq)" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5421 | by (meson finite_UN qq \<open>finite r\<close> tagged_division_of_finite) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5422 | qed | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5423 | moreover have "content M *\<^sub>R f x = 0" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5424 | if x: "(x,M) \<in> qq K" "(x,M) \<in> qq L" and KL: "qq K \<noteq> qq L" and r: "K \<in> r" "L \<in> r" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5425 | for x M K L | 
| 53638 | 5426 | proof - | 
| 5427 | note kl = tagged_division_ofD(3,4)[OF qq[THEN conjunct1]] | |
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5428 | obtain u v where uv: "M = cbox u v" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5429 | using \<open>(x, M) \<in> qq L\<close> \<open>L \<in> r\<close> kl(2) by blast | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5430 |     have empty: "interior (K \<inter> L) = {}"
 | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5431 | by (metis DiffD1 interior_Int q'(5) r_def KL r) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5432 |     have "interior M = {}"
 | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5433 | by (metis (no_types, lifting) Int_assoc empty inf.absorb_iff2 interior_Int kl(1) subset_empty x r) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5434 | then show "content M *\<^sub>R f x = 0" | 
| 53638 | 5435 | unfolding uv content_eq_0_interior[symmetric] | 
| 5436 | by auto | |
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5437 | qed | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5438 | ultimately have "norm (?SUM p + sum ?SUM (qq ` r) - integral (cbox a b) f) < e" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5439 | apply (subst (asm) sum.Union_comp) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5440 | using qq by (force simp: split_paired_all)+ | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5441 | moreover have "content M *\<^sub>R f x = 0" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5442 | if "K \<in> r" "L \<in> r" "K \<noteq> L" "qq K = qq L" "(x, M) \<in> qq K" for K L x M | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5443 | using tagged_division_ofD(6) qq that by (metis (no_types, lifting)) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5444 | ultimately have less_e: "norm (?SUM p + sum (?SUM \<circ> qq) r - integral (cbox a b) f) < e" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5445 | apply (subst (asm) sum.reindex_nontrivial [OF \<open>finite r\<close>]) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5446 | apply (auto simp: split_paired_all sum.neutral) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5447 | done | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5448 | have norm_le: "norm (cp - ip) \<le> e + k" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5449 | if "norm ((cp + cr) - i) < e" "norm (cr - ir) < k" "ip + ir = i" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5450 | for ir ip i cr cp::'a | 
| 53638 | 5451 | proof - | 
| 61165 | 5452 | from that show ?thesis | 
| 53638 | 5453 | using norm_triangle_le[of "cp + cr - i" "- (cr - ir)"] | 
| 61165 | 5454 | unfolding that(3)[symmetric] norm_minus_cancel | 
| 53638 | 5455 | by (auto simp add: algebra_simps) | 
| 5456 | qed | |
| 53399 | 5457 | |
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5458 | have "?lhs = norm (?SUM p - (\<Sum>(x, k)\<in>p. integral k f))" | 
| 64267 | 5459 | unfolding split_def sum_subtractf .. | 
| 53638 | 5460 | also have "\<dots> \<le> e + k" | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5461 | proof (rule norm_le[OF less_e]) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5462 | have lessk: "k * real (card r) / (1 + real (card r)) < k" | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5463 | using \<open>k>0\<close> by (auto simp add: field_simps) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5464 | have "norm (sum (?SUM \<circ> qq) r - (\<Sum>k\<in>r. integral k f)) \<le> (\<Sum>x\<in>r. k / (real (card r) + 1))" | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5465 | unfolding sum_subtractf[symmetric] by (force dest: qq intro!: sum_norm_le) | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5466 | also have "... < k" | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5467 | by (simp add: lessk add.commute mult.commute) | 
| 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5468 | finally show "norm (sum (?SUM \<circ> qq) r - (\<Sum>k\<in>r. integral k f)) < k" . | 
| 61167 | 5469 | next | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5470 | from q(1) have [simp]: "snd ` p \<union> q = q" by auto | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5471 | have "integral l f = 0" | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5472 | if inp: "(x, l) \<in> p" "(y, m) \<in> p" and ne: "(x, l) \<noteq> (y, m)" and "l = m" for x l y m | 
| 53638 | 5473 | proof - | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5474 | obtain u v where uv: "l = cbox u v" | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5475 | using inp p'(4) by blast | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5476 | have "content (cbox u v) = 0" | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5477 | unfolding content_eq_0_interior using that p(1) uv by auto | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5478 | then show ?thesis | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5479 | using uv by blast | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5480 | qed | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5481 | then have "(\<Sum>(x, k)\<in>p. integral k f) = (\<Sum>k\<in>snd ` p. integral k f)" | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5482 | apply (subst sum.reindex_nontrivial [OF \<open>finite p\<close>]) | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5483 | unfolding split_paired_all split_def by auto | 
| 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5484 | then show "(\<Sum>(x, k)\<in>p. integral k f) + (\<Sum>k\<in>r. integral k f) = integral (cbox a b) f" | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5485 | unfolding integral_combine_division_topdown[OF intf qdiv] r_def | 
| 66513 
ca8b18baf0e0
unscrambling esp of Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66512diff
changeset | 5486 | using q'(1) p'(1) sum.union_disjoint [of "snd ` p" "q - snd ` p", symmetric] | 
| 57418 | 5487 | by simp | 
| 53638 | 5488 | qed | 
| 66518 
5e65236e95aa
unscrambled Henstock_lemma_part1
 paulson <lp15@cam.ac.uk> parents: 
66513diff
changeset | 5489 | finally show "?lhs \<le> e + k" . | 
| 53638 | 5490 | qed | 
| 5491 | ||
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5492 | lemma Henstock_lemma_part2: | 
| 56188 | 5493 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 66437 | 5494 | assumes fed: "f integrable_on cbox a b" "e > 0" "gauge d" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5495 | and less_e: "\<And>\<D>. \<lbrakk>\<D> tagged_division_of (cbox a b); d fine \<D>\<rbrakk> \<Longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5496 | norm (sum (\<lambda>(x,k). content k *\<^sub>R f x) \<D> - integral (cbox a b) f) < e" | 
| 66437 | 5497 | and tag: "p tagged_partial_division_of (cbox a b)" | 
| 53638 | 5498 | and "d fine p" | 
| 64267 | 5499 |   shows "sum (\<lambda>(x,k). norm (content k *\<^sub>R f x - integral k f)) p \<le> 2 * real (DIM('n)) * e"
 | 
| 66437 | 5500 | proof - | 
| 5501 | have "finite p" | |
| 5502 | using tag by blast | |
| 5503 | then show ?thesis | |
| 5504 | unfolding split_def | |
| 5505 | proof (rule sum_norm_allsubsets_bound) | |
| 5506 | fix Q | |
| 5507 | assume Q: "Q \<subseteq> p" | |
| 5508 | then have fine: "d fine Q" | |
| 5509 | by (simp add: \<open>d fine p\<close> fine_subset) | |
| 5510 | show "norm (\<Sum>x\<in>Q. content (snd x) *\<^sub>R f (fst x) - integral (snd x) f) \<le> e" | |
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5511 | apply (rule Henstock_lemma_part1[OF fed less_e, unfolded split_def]) | 
| 66437 | 5512 | using Q tag tagged_partial_division_subset apply (force simp add: fine)+ | 
| 5513 | done | |
| 5514 | qed | |
| 5515 | qed | |
| 53638 | 5516 | |
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5517 | lemma Henstock_lemma: | 
| 56188 | 5518 | fixes f :: "'m::euclidean_space \<Rightarrow> 'n::euclidean_space" | 
| 66437 | 5519 | assumes intf: "f integrable_on cbox a b" | 
| 53638 | 5520 | and "e > 0" | 
| 66437 | 5521 | obtains \<gamma> where "gauge \<gamma>" | 
| 5522 | and "\<And>p. \<lbrakk>p tagged_partial_division_of (cbox a b); \<gamma> fine p\<rbrakk> \<Longrightarrow> | |
| 5523 | sum (\<lambda>(x,k). norm(content k *\<^sub>R f x - integral k f)) p < e" | |
| 53638 | 5524 | proof - | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5525 |   have *: "e/(2 * (real DIM('n) + 1)) > 0" using \<open>e > 0\<close> by simp
 | 
| 66437 | 5526 | with integrable_integral[OF intf, unfolded has_integral] | 
| 5527 | obtain \<gamma> where "gauge \<gamma>" | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5528 | and \<gamma>: "\<And>\<D>. \<lbrakk>\<D> tagged_division_of cbox a b; \<gamma> fine \<D>\<rbrakk> \<Longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5529 | norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f x) - integral (cbox a b) f) | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5530 |          < e/(2 * (real DIM('n) + 1))"
 | 
| 66437 | 5531 | by metis | 
| 53638 | 5532 | show thesis | 
| 66437 | 5533 | proof (rule that [OF \<open>gauge \<gamma>\<close>]) | 
| 5534 | fix p | |
| 5535 | assume p: "p tagged_partial_division_of cbox a b" "\<gamma> fine p" | |
| 5536 | have "(\<Sum>(x,K)\<in>p. norm (content K *\<^sub>R f x - integral K f)) | |
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5537 |           \<le> 2 * real DIM('n) * (e/(2 * (real DIM('n) + 1)))"
 | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5538 | using Henstock_lemma_part2[OF intf * \<open>gauge \<gamma>\<close> \<gamma> p] by metis | 
| 66437 | 5539 | also have "... < e" | 
| 5540 | using \<open>e > 0\<close> by (auto simp add: field_simps) | |
| 5541 | finally | |
| 5542 | show "(\<Sum>(x,K)\<in>p. norm (content K *\<^sub>R f x - integral K f)) < e" . | |
| 53638 | 5543 | qed | 
| 5544 | qed | |
| 5545 | ||
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 5546 | |
| 60420 | 5547 | subsection \<open>Monotone convergence (bounded interval first)\<close> | 
| 53638 | 5548 | |
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5549 | lemma bounded_increasing_convergent: | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5550 | fixes f :: "nat \<Rightarrow> real" | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5551 | shows "\<lbrakk>bounded (range f); \<And>n. f n \<le> f (Suc n)\<rbrakk> \<Longrightarrow> \<exists>l. f \<longlonglongrightarrow> l" | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5552 | using Bseq_mono_convergent[of f] incseq_Suc_iff[of f] | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5553 | by (auto simp: image_def Bseq_eq_bounded convergent_def incseq_def) | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5554 | |
| 53638 | 5555 | lemma monotone_convergence_interval: | 
| 56188 | 5556 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5557 | assumes intf: "\<And>k. (f k) integrable_on cbox a b" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5558 | and le: "\<And>k x. x \<in> cbox a b \<Longrightarrow> (f k x) \<le> f (Suc k) x" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5559 | and fg: "\<And>x. x \<in> cbox a b \<Longrightarrow> ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5560 | and bou: "bounded (range (\<lambda>k. integral (cbox a b) (f k)))" | 
| 61973 | 5561 | shows "g integrable_on cbox a b \<and> ((\<lambda>k. integral (cbox a b) (f k)) \<longlongrightarrow> integral (cbox a b) g) sequentially" | 
| 56188 | 5562 | proof (cases "content (cbox a b) = 0") | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5563 | case True then show ?thesis | 
| 53638 | 5564 | by auto | 
| 5565 | next | |
| 5566 | case False | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5567 | have fg1: "(f k x) \<le> (g x)" if x: "x \<in> cbox a b" for x k | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5568 | proof - | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5569 | have "\<forall>\<^sub>F j in sequentially. f k x \<le> f j x" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5570 | apply (rule eventually_sequentiallyI [of k]) | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5571 | using le x apply (force intro: transitive_stepwise_le) | 
| 53638 | 5572 | done | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5573 | then show "f k x \<le> g x" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5574 | using tendsto_lowerbound [OF fg] x trivial_limit_sequentially by blast | 
| 53638 | 5575 | qed | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5576 | have int_inc: "\<And>n. integral (cbox a b) (f n) \<le> integral (cbox a b) (f (Suc n))" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5577 | by (metis integral_le intf le) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5578 | then obtain i where i: "(\<lambda>k. integral (cbox a b) (f k)) \<longlonglongrightarrow> i" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5579 | using bounded_increasing_convergent bou by blast | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5580 | have "\<And>k. \<forall>\<^sub>F x in sequentially. integral (cbox a b) (f k) \<le> integral (cbox a b) (f x)" | 
| 53638 | 5581 | unfolding eventually_sequentially | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5582 | by (force intro: transitive_stepwise_le int_inc) | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5583 | then have i': "\<And>k. (integral(cbox a b) (f k)) \<le> i" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5584 | using tendsto_le [OF trivial_limit_sequentially i] by blast | 
| 56188 | 5585 | have "(g has_integral i) (cbox a b)" | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5586 | unfolding has_integral real_norm_def | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5587 | proof clarify | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5588 | fix e::real | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5589 | assume e: "e > 0" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5590 | have "\<And>k. (\<exists>\<gamma>. gauge \<gamma> \<and> (\<forall>\<D>. \<D> tagged_division_of (cbox a b) \<and> \<gamma> fine \<D> \<longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5591 | abs ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f k x) - integral (cbox a b) (f k)) < e/2 ^ (k + 2)))" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5592 | using intf e by (auto simp: has_integral_integral has_integral) | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5593 | then obtain c where c: "\<And>x. gauge (c x)" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5594 | "\<And>x \<D>. \<lbrakk>\<D> tagged_division_of cbox a b; c x fine \<D>\<rbrakk> \<Longrightarrow> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5595 | abs ((\<Sum>(u,K)\<in>\<D>. content K *\<^sub>R f x u) - integral (cbox a b) (f x)) | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5596 | < e/2 ^ (x + 2)" | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5597 | by metis | 
| 36243 
027ae62681be
Translated remaining theorems about integration from HOL light.
 himmelma parents: 
36081diff
changeset | 5598 | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5599 | have "\<exists>r. \<forall>k\<ge>r. 0 \<le> i - (integral (cbox a b) (f k)) \<and> i - (integral (cbox a b) (f k)) < e/4" | 
| 53638 | 5600 | proof - | 
| 5601 | have "e/4 > 0" | |
| 5602 | using e by auto | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5603 | show ?thesis | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5604 | using LIMSEQ_D [OF i \<open>e/4 > 0\<close>] i' by auto | 
| 53638 | 5605 | qed | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5606 | then obtain r where r: "\<And>k. r \<le> k \<Longrightarrow> 0 \<le> i - integral (cbox a b) (f k)" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5607 | "\<And>k. r \<le> k \<Longrightarrow> i - integral (cbox a b) (f k) < e/4" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5608 | by metis | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5609 | have "\<exists>n\<ge>r. \<forall>k\<ge>n. 0 \<le> (g x) - (f k x) \<and> (g x) - (f k x) < e/(4 * content(cbox a b))" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5610 | if "x \<in> cbox a b" for x | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5611 | proof - | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5612 | have "e/(4 * content (cbox a b)) > 0" | 
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 5613 | by (simp add: False content_lt_nz e) | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5614 | with fg that LIMSEQ_D | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5615 | obtain N where "\<forall>n\<ge>N. norm (f n x - g x) < e/(4 * content (cbox a b))" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5616 | by metis | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5617 | then show "\<exists>n\<ge>r. | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5618 | \<forall>k\<ge>n. | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5619 | 0 \<le> g x - f k x \<and> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5620 | g x - f k x | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5621 | < e/(4 * content (cbox a b))" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5622 | apply (rule_tac x="N + r" in exI) | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5623 | using fg1[OF that] apply (auto simp add: field_simps) | 
| 53638 | 5624 | done | 
| 5625 | qed | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5626 | then obtain m where r_le_m: "\<And>x. x \<in> cbox a b \<Longrightarrow> r \<le> m x" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5627 | and m: "\<And>x k. \<lbrakk>x \<in> cbox a b; m x \<le> k\<rbrakk> | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5628 | \<Longrightarrow> 0 \<le> g x - f k x \<and> g x - f k x < e/(4 * content (cbox a b))" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5629 | by metis | 
| 63040 | 5630 | define d where "d x = c (m x) x" for x | 
| 66409 | 5631 | show "\<exists>\<gamma>. gauge \<gamma> \<and> | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5632 | (\<forall>\<D>. \<D> tagged_division_of cbox a b \<and> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5633 | \<gamma> fine \<D> \<longrightarrow> abs ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - i) < e)" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5634 | proof (rule exI, safe) | 
| 53638 | 5635 | show "gauge d" | 
| 5636 | using c(1) unfolding gauge_def d_def by auto | |
| 5637 | next | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5638 | fix \<D> | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5639 | assume ptag: "\<D> tagged_division_of (cbox a b)" and "d fine \<D>" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5640 | note p'=tagged_division_ofD[OF ptag] | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5641 | obtain s where s: "\<And>x. x \<in> \<D> \<Longrightarrow> m (fst x) \<le> s" | 
| 41851 | 5642 | by (metis finite_imageI finite_nat_set_iff_bounded_le p'(1) rev_image_eqI) | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5643 | have *: "\<bar>a - d\<bar> < e" if "\<bar>a - b\<bar> \<le> e/4" "\<bar>b - c\<bar> < e/2" "\<bar>c - d\<bar> < e/4" for a b c d | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5644 | using that norm_triangle_lt[of "a - b" "b - c" "3* e/4"] | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5645 | norm_triangle_lt[of "a - b + (b - c)" "c - d" e] | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5646 | by (auto simp add: algebra_simps) | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5647 | show "\<bar>(\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R g x) - i\<bar> < e" | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5648 | proof (rule *) | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5649 | have "\<bar>(\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f (m x) x)\<bar> | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5650 | \<le> (\<Sum>i\<in>\<D>. \<bar>(case i of (x, K) \<Rightarrow> content K *\<^sub>R g x) - (case i of (x, K) \<Rightarrow> content K *\<^sub>R f (m x) x)\<bar>)" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5651 | by (metis (mono_tags) sum_subtractf sum_abs) | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5652 | also have "... \<le> (\<Sum>(x, k)\<in>\<D>. content k * (e/(4 * content (cbox a b))))" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5653 | proof (rule sum_mono, simp add: split_paired_all) | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5654 | fix x K | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5655 | assume xk: "(x,K) \<in> \<D>" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5656 | with ptag have x: "x \<in> cbox a b" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5657 | by blast | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5658 | then have "abs (content K * (g x - f (m x) x)) \<le> content K * (e/(4 * content (cbox a b)))" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5659 | by (metis m[OF x] mult_nonneg_nonneg abs_of_nonneg less_eq_real_def measure_nonneg mult_left_mono order_refl) | 
| 66505 | 5660 | then show "\<bar>content K * g x - content K * f (m x) x\<bar> \<le> content K * e/(4 * content (cbox a b))" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5661 | by (simp add: algebra_simps) | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5662 | qed | 
| 66505 | 5663 | also have "... = (e/(4 * content (cbox a b))) * (\<Sum>(x, k)\<in>\<D>. content k)" | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5664 | by (simp add: sum_distrib_left sum_divide_distrib split_def mult.commute) | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5665 | also have "... \<le> e/4" | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5666 | by (metis False additive_content_tagged_division [OF ptag] nonzero_mult_divide_mult_cancel_right order_refl times_divide_eq_left) | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5667 | finally show "\<bar>(\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R g x) - (\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f (m x) x)\<bar> \<le> e/4" . | 
| 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5668 | |
| 53638 | 5669 | next | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5670 | have "norm ((\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f (m x) x) - (\<Sum>(x,K)\<in>\<D>. integral K (f (m x)))) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5671 |               \<le> norm (\<Sum>j = 0..s. \<Sum>(x,K)\<in>{xk \<in> \<D>. m (fst xk) = j}. content K *\<^sub>R f (m x) x - integral K (f (m x)))"
 | 
| 64267 | 5672 | apply (subst sum_group) | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5673 | using s by (auto simp: sum_subtractf split_def p'(1)) | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5674 | also have "\<dots> < e/2" | 
| 53638 | 5675 | proof - | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5676 |           have "norm (\<Sum>j = 0..s. \<Sum>(x, k)\<in>{xk \<in> \<D>. m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x)))
 | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5677 | \<le> (\<Sum>i = 0..s. e/2 ^ (i + 2))" | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5678 | proof (rule sum_norm_le) | 
| 53638 | 5679 | fix t | 
| 5680 |             assume "t \<in> {0..s}"
 | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5681 |             have "norm (\<Sum>(x,k)\<in>{xk \<in> \<D>. m (fst xk) = t}. content k *\<^sub>R f (m x) x - integral k (f (m x))) =
 | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5682 |                   norm (\<Sum>(x,k)\<in>{xk \<in> \<D>. m (fst xk) = t}. content k *\<^sub>R f t x - integral k (f t))"
 | 
| 66437 | 5683 | by (force intro!: sum.cong arg_cong[where f=norm]) | 
| 5684 | also have "... \<le> e/2 ^ (t + 2)" | |
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5685 | proof (rule Henstock_lemma_part1 [OF intf]) | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5686 |               show "{xk \<in> \<D>. m (fst xk) = t} tagged_partial_division_of cbox a b"
 | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5687 | apply (rule tagged_partial_division_subset[of \<D>]) | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5688 | using ptag by (auto simp: tagged_division_of_def) | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5689 |               show "c t fine {xk \<in> \<D>. m (fst xk) = t}"
 | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5690 | using \<open>d fine \<D>\<close> by (auto simp: fine_def d_def) | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5691 | qed (use c e in auto) | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5692 |             finally show "norm (\<Sum>(x,K)\<in>{xk \<in> \<D>. m (fst xk) = t}. content K *\<^sub>R f (m x) x -
 | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5693 | integral K (f (m x))) \<le> e/2 ^ (t + 2)" . | 
| 53638 | 5694 | qed | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5695 | also have "... = (e/2/2) * (\<Sum>i = 0..s. (1/2) ^ i)" | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5696 | by (simp add: sum_distrib_left field_simps) | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5697 | also have "\<dots> < e/2" | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5698 | by (simp add: sum_gp mult_strict_left_mono[OF _ e]) | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5699 |           finally show "norm (\<Sum>j = 0..s. \<Sum>(x, k)\<in>{xk \<in> \<D>.
 | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5700 | m (fst xk) = j}. content k *\<^sub>R f (m x) x - integral k (f (m x))) < e/2" . | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5701 | qed | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5702 | finally show "\<bar>(\<Sum>(x,K)\<in>\<D>. content K *\<^sub>R f (m x) x) - (\<Sum>(x,K)\<in>\<D>. integral K (f (m x)))\<bar> < e/2" | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5703 | by simp | 
| 53638 | 5704 | next | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5705 | have comb: "integral (cbox a b) (f y) = (\<Sum>(x, k)\<in>\<D>. integral k (f y))" for y | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5706 | using integral_combine_tagged_division_topdown[OF intf ptag] by metis | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5707 | have f_le: "\<And>y m n. \<lbrakk>y \<in> cbox a b; n\<ge>m\<rbrakk> \<Longrightarrow> f m y \<le> f n y" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5708 | using le by (auto intro: transitive_stepwise_le) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5709 | have "(\<Sum>(x, k)\<in>\<D>. integral k (f r)) \<le> (\<Sum>(x, K)\<in>\<D>. integral K (f (m x)))" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5710 | proof (rule sum_mono, simp add: split_paired_all) | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5711 | fix x K | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5712 | assume xK: "(x, K) \<in> \<D>" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5713 | show "integral K (f r) \<le> integral K (f (m x))" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5714 | proof (rule integral_le) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5715 | show "f r integrable_on K" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5716 | by (metis integrable_on_subcbox intf p'(3) p'(4) xK) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5717 | show "f (m x) integrable_on K" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5718 | by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5719 | show "f r y \<le> f (m x) y" if "y \<in> K" for y | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5720 | using that r_le_m[of x] p'(2-3)[OF xK] f_le by auto | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5721 | qed | 
| 53638 | 5722 | qed | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5723 | moreover have "(\<Sum>(x, K)\<in>\<D>. integral K (f (m x))) \<le> (\<Sum>(x, k)\<in>\<D>. integral k (f s))" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5724 | proof (rule sum_mono, simp add: split_paired_all) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5725 | fix x K | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5726 | assume xK: "(x, K) \<in> \<D>" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5727 | show "integral K (f (m x)) \<le> integral K (f s)" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5728 | proof (rule integral_le) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5729 | show "f (m x) integrable_on K" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5730 | by (metis elementary_interval integrable_on_subdivision intf p'(3) p'(4) xK) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5731 | show "f s integrable_on K" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5732 | by (metis integrable_on_subcbox intf p'(3) p'(4) xK) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5733 | show "f (m x) y \<le> f s y" if "y \<in> K" for y | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5734 | using that s xK f_le p'(3) by fastforce | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5735 | qed | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5736 | qed | 
| 66505 | 5737 | moreover have "0 \<le> i - integral (cbox a b) (f r)" "i - integral (cbox a b) (f r) < e/4" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5738 | using r by auto | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5739 | ultimately show "\<bar>(\<Sum>(x,K)\<in>\<D>. integral K (f (m x))) - i\<bar> < e/4" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5740 | using comb i'[of s] by auto | 
| 53638 | 5741 | qed | 
| 5742 | qed | |
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5743 | qed | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5744 | with i integral_unique show ?thesis | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5745 | by blast | 
| 53638 | 5746 | qed | 
| 5747 | ||
| 5748 | lemma monotone_convergence_increasing: | |
| 56188 | 5749 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real" | 
| 66420 | 5750 | assumes int_f: "\<And>k. (f k) integrable_on S" | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5751 | and "\<And>k x. x \<in> S \<Longrightarrow> (f k x) \<le> (f (Suc k) x)" | 
| 66420 | 5752 | and fg: "\<And>x. x \<in> S \<Longrightarrow> ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially" | 
| 5753 | and bou: "bounded (range (\<lambda>k. integral S (f k)))" | |
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5754 | shows "g integrable_on S \<and> ((\<lambda>k. integral S (f k)) \<longlongrightarrow> integral S g) sequentially" | 
| 61973 | 5755 | proof - | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5756 | have lem: "g integrable_on S \<and> ((\<lambda>k. integral S (f k)) \<longlongrightarrow> integral S g) sequentially" | 
| 66409 | 5757 | if f0: "\<And>k x. x \<in> S \<Longrightarrow> 0 \<le> f k x" | 
| 5758 | and int_f: "\<And>k. (f k) integrable_on S" | |
| 66420 | 5759 | and le: "\<And>k x. x \<in> S \<Longrightarrow> f k x \<le> f (Suc k) x" | 
| 66409 | 5760 | and lim: "\<And>x. x \<in> S \<Longrightarrow> ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially" | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5761 | and bou: "bounded (range(\<lambda>k. integral S (f k)))" | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5762 | for f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real" and g S | 
| 53638 | 5763 | proof - | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5764 | have fg: "(f k x) \<le> (g x)" if "x \<in> S" for x k | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5765 | apply (rule tendsto_lowerbound [OF lim [OF that]]) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5766 | apply (rule eventually_sequentiallyI [of k]) | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5767 | using le by (force intro: transitive_stepwise_le that)+ | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5768 | obtain i where i: "(\<lambda>k. integral S (f k)) \<longlonglongrightarrow> i" | 
| 66420 | 5769 | using bounded_increasing_convergent [OF bou] le int_f integral_le by blast | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5770 | have i': "(integral S (f k)) \<le> i" for k | 
| 66409 | 5771 | proof - | 
| 66519 | 5772 | have "\<And>k. \<And>x. x \<in> S \<Longrightarrow> \<forall>n\<ge>k. f k x \<le> f n x" | 
| 66409 | 5773 | using le by (force intro: transitive_stepwise_le) | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5774 | then show ?thesis | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5775 | using tendsto_lowerbound [OF i eventually_sequentiallyI trivial_limit_sequentially] | 
| 66497 
18a6478a574c
More tidying, and renaming of theorems
 paulson <lp15@cam.ac.uk> parents: 
66495diff
changeset | 5776 | by (meson int_f integral_le) | 
| 66409 | 5777 | qed | 
| 5778 | ||
| 66420 | 5779 | let ?f = "(\<lambda>k x. if x \<in> S then f k x else 0)" | 
| 5780 | let ?g = "(\<lambda>x. if x \<in> S then g x else 0)" | |
| 5781 | have int: "?f k integrable_on cbox a b" for a b k | |
| 66409 | 5782 | by (simp add: int_f integrable_altD(1)) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5783 | have int': "\<And>k a b. f k integrable_on cbox a b \<inter> S" | 
| 66409 | 5784 | using int by (simp add: Int_commute integrable_restrict_Int) | 
| 66420 | 5785 | have g: "?g integrable_on cbox a b \<and> | 
| 5786 | (\<lambda>k. integral (cbox a b) (?f k)) \<longlonglongrightarrow> integral (cbox a b) ?g" for a b | |
| 5787 | proof (rule monotone_convergence_interval) | |
| 5788 | have "norm (integral (cbox a b) (?f k)) \<le> norm (integral S (f k))" for k | |
| 66409 | 5789 | proof - | 
| 66420 | 5790 | have "0 \<le> integral (cbox a b) (?f k)" | 
| 66409 | 5791 | by (metis (no_types) integral_nonneg Int_iff f0 inf_commute integral_restrict_Int int') | 
| 5792 | moreover have "0 \<le> integral S (f k)" | |
| 5793 | by (simp add: integral_nonneg f0 int_f) | |
| 5794 | moreover have "integral (S \<inter> cbox a b) (f k) \<le> integral S (f k)" | |
| 5795 | by (metis f0 inf_commute int' int_f integral_subset_le le_inf_iff order_refl) | |
| 5796 | ultimately show ?thesis | |
| 5797 | by (simp add: integral_restrict_Int) | |
| 5798 | qed | |
| 5799 | moreover obtain B where "\<And>x. x \<in> range (\<lambda>k. integral S (f k)) \<Longrightarrow> norm x \<le> B" | |
| 5800 | using bou unfolding bounded_iff by blast | |
| 66420 | 5801 | ultimately show "bounded (range (\<lambda>k. integral (cbox a b) (?f k)))" | 
| 66409 | 5802 | unfolding bounded_iff by (blast intro: order_trans) | 
| 66420 | 5803 | qed (use int le lim in auto) | 
| 5804 | moreover have "\<exists>B>0. \<forall>a b. ball 0 B \<subseteq> cbox a b \<longrightarrow> norm (integral (cbox a b) ?g - i) < e" | |
| 5805 | if "0 < e" for e | |
| 5806 | proof - | |
| 5807 | have "e/4>0" | |
| 5808 | using that by auto | |
| 5809 | with LIMSEQ_D [OF i] obtain N where N: "\<And>n. n \<ge> N \<Longrightarrow> norm (integral S (f n) - i) < e/4" | |
| 5810 | by metis | |
| 5811 | with int_f[of N, unfolded has_integral_integral has_integral_alt'[of "f N"]] | |
| 5812 | obtain B where "0 < B" and B: | |
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5813 | "\<And>a b. ball 0 B \<subseteq> cbox a b \<Longrightarrow> norm (integral (cbox a b) (?f N) - integral S (f N)) < e/4" | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5814 | by (meson \<open>0 < e/4\<close>) | 
| 66420 | 5815 | have "norm (integral (cbox a b) ?g - i) < e" if ab: "ball 0 B \<subseteq> cbox a b" for a b | 
| 53638 | 5816 | proof - | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5817 | obtain M where M: "\<And>n. n \<ge> M \<Longrightarrow> abs (integral (cbox a b) (?f n) - integral (cbox a b) ?g) < e/2" | 
| 66420 | 5818 | using \<open>e > 0\<close> g by (fastforce simp add: dest!: LIMSEQ_D [where r = "e/2"]) | 
| 5819 | have *: "\<And>\<alpha> \<beta> g. \<lbrakk>\<bar>\<alpha> - i\<bar> < e/2; \<bar>\<beta> - g\<bar> < e/2; \<alpha> \<le> \<beta>; \<beta> \<le> i\<rbrakk> \<Longrightarrow> \<bar>g - i\<bar> < e" | |
| 53638 | 5820 | unfolding real_inner_1_right by arith | 
| 66420 | 5821 | show "norm (integral (cbox a b) ?g - i) < e" | 
| 53638 | 5822 | unfolding real_norm_def | 
| 66420 | 5823 | proof (rule *) | 
| 5824 | show "\<bar>integral (cbox a b) (?f N) - i\<bar> < e/2" | |
| 5825 | proof (rule abs_triangle_half_l) | |
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 5826 | show "\<bar>integral (cbox a b) (?f N) - integral S (f N)\<bar> < e/2/2" | 
| 66420 | 5827 | using B[OF ab] by simp | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 5828 | show "abs (i - integral S (f N)) < e/2/2" | 
| 66420 | 5829 | using N by (simp add: abs_minus_commute) | 
| 5830 | qed | |
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 5831 | show "\<bar>integral (cbox a b) (?f (M + N)) - integral (cbox a b) ?g\<bar> < e/2" | 
| 66420 | 5832 | by (metis le_add1 M[of "M + N"]) | 
| 5833 | show "integral (cbox a b) (?f N) \<le> integral (cbox a b) (?f (M + N))" | |
| 5834 | proof (intro ballI integral_le[OF int int]) | |
| 5835 | fix x assume "x \<in> cbox a b" | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5836 | have "(f m x) \<le> (f n x)" if "x \<in> S" "n \<ge> m" for m n | 
| 66420 | 5837 | apply (rule transitive_stepwise_le [OF \<open>n \<ge> m\<close> order_refl]) | 
| 5838 | using dual_order.trans apply blast | |
| 5839 | by (simp add: le \<open>x \<in> S\<close>) | |
| 5840 | then show "(?f N)x \<le> (?f (M+N))x" | |
| 5841 | by auto | |
| 5842 | qed | |
| 66409 | 5843 | have "integral (cbox a b \<inter> S) (f (M + N)) \<le> integral S (f (M + N))" | 
| 5844 | by (metis Int_lower1 f0 inf_commute int' int_f integral_subset_le) | |
| 66420 | 5845 | then have "integral (cbox a b) (?f (M + N)) \<le> integral S (f (M + N))" | 
| 5846 | by (metis (no_types) inf_commute integral_restrict_Int) | |
| 5847 | also have "... \<le> i" | |
| 5848 | using i'[of "M + N"] by auto | |
| 5849 | finally show "integral (cbox a b) (?f (M + N)) \<le> i" . | |
| 53638 | 5850 | qed | 
| 5851 | qed | |
| 66420 | 5852 | then show ?thesis | 
| 5853 | using \<open>0 < B\<close> by blast | |
| 53638 | 5854 | qed | 
| 66420 | 5855 | ultimately have "(g has_integral i) S" | 
| 5856 | unfolding has_integral_alt' by auto | |
| 61165 | 5857 | then show ?thesis | 
| 66409 | 5858 | using has_integral_integrable_integral i integral_unique by metis | 
| 53638 | 5859 | qed | 
| 5860 | ||
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5861 | have sub: "\<And>k. integral S (\<lambda>x. f k x - f 0 x) = integral S (f k) - integral S (f 0)" | 
| 66420 | 5862 | by (simp add: integral_diff int_f) | 
| 66409 | 5863 | have *: "\<And>x m n. x \<in> S \<Longrightarrow> n\<ge>m \<Longrightarrow> f m x \<le> f n x" | 
| 66193 | 5864 | using assms(2) by (force intro: transitive_stepwise_le) | 
| 66420 | 5865 | have gf: "(\<lambda>x. g x - f 0 x) integrable_on S \<and> ((\<lambda>k. integral S (\<lambda>x. f (Suc k) x - f 0 x)) \<longlongrightarrow> | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5866 | integral S (\<lambda>x. g x - f 0 x)) sequentially" | 
| 66420 | 5867 | proof (rule lem) | 
| 5868 | show "\<And>k. (\<lambda>x. f (Suc k) x - f 0 x) integrable_on S" | |
| 5869 | by (simp add: integrable_diff int_f) | |
| 5870 | show "(\<lambda>k. f (Suc k) x - f 0 x) \<longlonglongrightarrow> g x - f 0 x" if "x \<in> S" for x | |
| 5871 | proof - | |
| 5872 | have "(\<lambda>n. f (Suc n) x) \<longlonglongrightarrow> g x" | |
| 5873 | using LIMSEQ_ignore_initial_segment[OF fg[OF \<open>x \<in> S\<close>], of 1] by simp | |
| 5874 | then show ?thesis | |
| 5875 | by (simp add: tendsto_diff) | |
| 5876 | qed | |
| 5877 | show "bounded (range (\<lambda>k. integral S (\<lambda>x. f (Suc k) x - f 0 x)))" | |
| 5878 | proof - | |
| 5879 | obtain B where B: "\<And>k. norm (integral S (f k)) \<le> B" | |
| 5880 | using bou by (auto simp: bounded_iff) | |
| 5881 | then have "norm (integral S (\<lambda>x. f (Suc k) x - f 0 x)) | |
| 5882 | \<le> B + norm (integral S (f 0))" for k | |
| 5883 | unfolding sub by (meson add_le_cancel_right norm_triangle_le_diff) | |
| 5884 | then show ?thesis | |
| 5885 | unfolding bounded_iff by blast | |
| 5886 | qed | |
| 5887 | qed (use * in auto) | |
| 5888 | then have "(\<lambda>x. integral S (\<lambda>xa. f (Suc x) xa - f 0 xa) + integral S (f 0)) | |
| 5889 | \<longlonglongrightarrow> integral S (\<lambda>x. g x - f 0 x) + integral S (f 0)" | |
| 5890 | by (auto simp add: tendsto_add) | |
| 5891 | moreover have "(\<lambda>x. g x - f 0 x + f 0 x) integrable_on S" | |
| 5892 | using gf integrable_add int_f [of 0] by metis | |
| 5893 | ultimately show ?thesis | |
| 5894 | by (simp add: integral_diff int_f LIMSEQ_imp_Suc sub) | |
| 53638 | 5895 | qed | 
| 5896 | ||
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5897 | lemma has_integral_monotone_convergence_increasing: | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5898 | fixes f :: "nat \<Rightarrow> 'a::euclidean_space \<Rightarrow> real" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5899 | assumes f: "\<And>k. (f k has_integral x k) s" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5900 | assumes "\<And>k x. x \<in> s \<Longrightarrow> f k x \<le> f (Suc k) x" | 
| 61969 | 5901 | assumes "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>k. f k x) \<longlonglongrightarrow> g x" | 
| 5902 | assumes "x \<longlonglongrightarrow> x'" | |
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5903 | shows "(g has_integral x') s" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5904 | proof - | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5905 | have x_eq: "x = (\<lambda>i. integral s (f i))" | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5906 | by (simp add: integral_unique[OF f]) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5907 | then have x: "range(\<lambda>k. integral s (f k)) = range x" | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5908 | by auto | 
| 63540 | 5909 | have *: "g integrable_on s \<and> (\<lambda>k. integral s (f k)) \<longlonglongrightarrow> integral s g" | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5910 | proof (intro monotone_convergence_increasing allI ballI assms) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5911 | show "bounded (range(\<lambda>k. integral s (f k)))" | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5912 | using x convergent_imp_bounded assms by metis | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 5913 | qed (use f in auto) | 
| 63540 | 5914 | then have "integral s g = x'" | 
| 61969 | 5915 | by (intro LIMSEQ_unique[OF _ \<open>x \<longlonglongrightarrow> x'\<close>]) (simp add: x_eq) | 
| 63540 | 5916 | with * show ?thesis | 
| 57447 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5917 | by (simp add: has_integral_integral) | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5918 | qed | 
| 
87429bdecad5
import more stuff from the CLT proof; base the lborel measure on interval_measure; remove lebesgue measure
 hoelzl parents: 
57418diff
changeset | 5919 | |
| 53638 | 5920 | lemma monotone_convergence_decreasing: | 
| 56188 | 5921 | fixes f :: "nat \<Rightarrow> 'n::euclidean_space \<Rightarrow> real" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5922 | assumes intf: "\<And>k. (f k) integrable_on S" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5923 | and le: "\<And>k x. x \<in> S \<Longrightarrow> f (Suc k) x \<le> f k x" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5924 | and fg: "\<And>x. x \<in> S \<Longrightarrow> ((\<lambda>k. f k x) \<longlongrightarrow> g x) sequentially" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5925 | and bou: "bounded (range(\<lambda>k. integral S (f k)))" | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5926 | shows "g integrable_on S \<and> (\<lambda>k. integral S (f k)) \<longlonglongrightarrow> integral S g" | 
| 53638 | 5927 | proof - | 
| 67399 | 5928 | have *: "range(\<lambda>k. integral S (\<lambda>x. - f k x)) = ( *\<^sub>R) (- 1) ` (range(\<lambda>k. integral S (f k)))" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5929 | by force | 
| 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5930 | have "(\<lambda>x. - g x) integrable_on S \<and> (\<lambda>k. integral S (\<lambda>x. - f k x)) \<longlonglongrightarrow> integral S (\<lambda>x. - g x)" | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5931 | proof (rule monotone_convergence_increasing) | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5932 | show "\<And>k. (\<lambda>x. - f k x) integrable_on S" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5933 | by (blast intro: integrable_neg intf) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5934 | show "\<And>k x. x \<in> S \<Longrightarrow> - f k x \<le> - f (Suc k) x" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5935 | by (simp add: le) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5936 | show "\<And>x. x \<in> S \<Longrightarrow> (\<lambda>k. - f k x) \<longlonglongrightarrow> - g x" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5937 | by (simp add: fg tendsto_minus) | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5938 | show "bounded (range(\<lambda>k. integral S (\<lambda>x. - f k x)))" | 
| 66429 
beaeb40a1217
tackling another nightmare proof
 paulson <lp15@cam.ac.uk> parents: 
66422diff
changeset | 5939 | using "*" bou bounded_scaling by auto | 
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5940 | qed | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5941 | then show ?thesis | 
| 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 5942 | by (force dest: integrable_neg tendsto_minus) | 
| 53638 | 5943 | qed | 
| 5944 | ||
| 5945 | lemma integral_norm_bound_integral: | |
| 56188 | 5946 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5947 | assumes int_f: "f integrable_on S" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5948 | and int_g: "g integrable_on S" | 
| 66519 | 5949 | and le_g: "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> g x" | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5950 | shows "norm (integral S f) \<le> integral S g" | 
| 53638 | 5951 | proof - | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 5952 | have norm: "norm \<eta> \<le> y + e" | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5953 | if "norm \<zeta> \<le> x" and "\<bar>x - y\<bar> < e/2" and "norm (\<zeta> - \<eta>) < e/2" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5954 | for e x y and \<zeta> \<eta> :: 'a | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5955 | proof - | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5956 | have "norm (\<eta> - \<zeta>) < e/2" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5957 | by (metis norm_minus_commute that(3)) | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5958 | moreover have "x \<le> y + e/2" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5959 | using that(2) by linarith | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5960 | ultimately show ?thesis | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 5961 | using that(1) le_less_trans[OF norm_triangle_sub[of \<eta> \<zeta>]] by (auto simp: less_imp_le) | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 5962 | qed | 
| 61165 | 5963 | have lem: "norm (integral(cbox a b) f) \<le> integral (cbox a b) g" | 
| 66359 | 5964 | if f: "f integrable_on cbox a b" | 
| 5965 | and g: "g integrable_on cbox a b" | |
| 5966 | and nle: "\<And>x. x \<in> cbox a b \<Longrightarrow> norm (f x) \<le> g x" | |
| 61165 | 5967 | for f :: "'n \<Rightarrow> 'a" and g a b | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 5968 | proof (rule field_le_epsilon) | 
| 61165 | 5969 | fix e :: real | 
| 5970 | assume "e > 0" | |
| 66359 | 5971 | then have e: "e/2 > 0" | 
| 53638 | 5972 | by auto | 
| 66359 | 5973 | with integrable_integral[OF f,unfolded has_integral[of f]] | 
| 5974 | obtain \<gamma> where \<gamma>: "gauge \<gamma>" | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5975 | "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> \<gamma> fine \<D> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5976 | \<Longrightarrow> norm ((\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" | 
| 66359 | 5977 | by meson | 
| 5978 | moreover | |
| 5979 | from integrable_integral[OF g,unfolded has_integral[of g]] e | |
| 5980 | obtain \<delta> where \<delta>: "gauge \<delta>" | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5981 | "\<And>\<D>. \<D> tagged_division_of cbox a b \<and> \<delta> fine \<D> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5982 | \<Longrightarrow> norm ((\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R g x) - integral (cbox a b) g) < e/2" | 
| 66359 | 5983 | by meson | 
| 5984 | ultimately have "gauge (\<lambda>x. \<gamma> x \<inter> \<delta> x)" | |
| 5985 | using gauge_Int by blast | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5986 | with fine_division_exists obtain \<D> | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5987 | where p: "\<D> tagged_division_of cbox a b" "(\<lambda>x. \<gamma> x \<inter> \<delta> x) fine \<D>" | 
| 66359 | 5988 | by metis | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5989 | have "\<gamma> fine \<D>" "\<delta> fine \<D>" | 
| 66359 | 5990 | using fine_Int p(2) by blast+ | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 5991 | show "norm (integral (cbox a b) f) \<le> integral (cbox a b) g + e" | 
| 66359 | 5992 | proof (rule norm) | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5993 | have "norm (content K *\<^sub>R f x) \<le> content K *\<^sub>R g x" if "(x, K) \<in> \<D>" for x K | 
| 66359 | 5994 | proof- | 
| 5995 | have K: "x \<in> K" "K \<subseteq> cbox a b" | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5996 | using \<open>(x, K) \<in> \<D>\<close> p(1) by blast+ | 
| 66359 | 5997 | obtain u v where "K = cbox u v" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 5998 | using \<open>(x, K) \<in> \<D>\<close> p(1) by blast | 
| 66359 | 5999 | moreover have "content K * norm (f x) \<le> content K * g x" | 
| 6000 | by (metis K subsetD dual_order.antisym measure_nonneg mult_zero_left nle not_le real_mult_le_cancel_iff2) | |
| 6001 | then show ?thesis | |
| 6002 | by simp | |
| 6003 | qed | |
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 6004 | then show "norm (\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R f x) \<le> (\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R g x)" | 
| 66359 | 6005 | by (simp add: sum_norm_le split_def) | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 6006 | show "norm ((\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R f x) - integral (cbox a b) f) < e/2" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 6007 | using \<open>\<gamma> fine \<D>\<close> \<gamma> p(1) by simp | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 6008 | show "\<bar>(\<Sum>(x, k)\<in>\<D>. content k *\<^sub>R g x) - integral (cbox a b) g\<bar> < e/2" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 6009 | using \<open>\<delta> fine \<D>\<close> \<delta> p(1) by simp | 
| 66359 | 6010 | qed | 
| 53638 | 6011 | qed | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6012 | show ?thesis | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 6013 | proof (rule field_le_epsilon) | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6014 | fix e :: real | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6015 | assume "e > 0" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6016 | then have e: "e/2 > 0" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6017 | by auto | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6018 | let ?f = "(\<lambda>x. if x \<in> S then f x else 0)" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6019 | let ?g = "(\<lambda>x. if x \<in> S then g x else 0)" | 
| 66420 | 6020 | have f: "?f integrable_on cbox a b" and g: "?g integrable_on cbox a b" for a b | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6021 | using int_f int_g integrable_altD by auto | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6022 | obtain Bf where "0 < Bf" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6023 | and Bf: "\<And>a b. ball 0 Bf \<subseteq> cbox a b \<Longrightarrow> | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6024 | \<exists>z. (?f has_integral z) (cbox a b) \<and> norm (z - integral S f) < e/2" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6025 | using integrable_integral [OF int_f,unfolded has_integral'[of f]] e that by blast | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6026 | obtain Bg where "0 < Bg" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6027 | and Bg: "\<And>a b. ball 0 Bg \<subseteq> cbox a b \<Longrightarrow> | 
| 66420 | 6028 | \<exists>z. (?g has_integral z) (cbox a b) \<and> norm (z - integral S g) < e/2" | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6029 | using integrable_integral [OF int_g,unfolded has_integral'[of g]] e that by blast | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6030 | obtain a b::'n where ab: "ball 0 Bf \<union> ball 0 Bg \<subseteq> cbox a b" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6031 | using ball_max_Un bounded_subset_cbox[OF bounded_ball, of _ "max Bf Bg"] by blast | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6032 | have "ball 0 Bf \<subseteq> cbox a b" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6033 | using ab by auto | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6034 | with Bf obtain z where int_fz: "(?f has_integral z) (cbox a b)" and z: "norm (z - integral S f) < e/2" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6035 | by meson | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6036 | have "ball 0 Bg \<subseteq> cbox a b" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6037 | using ab by auto | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6038 | with Bg obtain w where int_gw: "(?g has_integral w) (cbox a b)" and w: "norm (w - integral S g) < e/2" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6039 | by meson | 
| 66503 
7685861f337d
more elimination of "guess", etc.
 paulson <lp15@cam.ac.uk> parents: 
66498diff
changeset | 6040 | show "norm (integral S f) \<le> integral S g + e" | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6041 | proof (rule norm) | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6042 | show "norm (integral (cbox a b) ?f) \<le> integral (cbox a b) ?g" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6043 | by (simp add: le_g lem[OF f g, of a b]) | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6044 | show "\<bar>integral (cbox a b) ?g - integral S g\<bar> < e/2" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6045 | using int_gw integral_unique w by auto | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6046 | show "norm (integral (cbox a b) ?f - integral S f) < e/2" | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6047 | using int_fz integral_unique z by blast | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6048 | qed | 
| 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6049 | qed | 
| 53638 | 6050 | qed | 
| 6051 | ||
| 66359 | 6052 | |
| 53638 | 6053 | lemma integral_norm_bound_integral_component: | 
| 56188 | 6054 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 6055 | fixes g :: "'n \<Rightarrow> 'b::euclidean_space" | |
| 66519 | 6056 | assumes f: "f integrable_on S" and g: "g integrable_on S" | 
| 6057 | and fg: "\<And>x. x \<in> S \<Longrightarrow> norm(f x) \<le> (g x)\<bullet>k" | |
| 6058 | shows "norm (integral S f) \<le> (integral S g)\<bullet>k" | |
| 53638 | 6059 | proof - | 
| 66519 | 6060 | have "norm (integral S f) \<le> integral S ((\<lambda>x. x \<bullet> k) \<circ> g)" | 
| 6061 | apply (rule integral_norm_bound_integral[OF f integrable_linear[OF g]]) | |
| 6062 | apply (simp add: bounded_linear_inner_left) | |
| 66703 | 6063 | apply (metis fg o_def) | 
| 53638 | 6064 | done | 
| 6065 | then show ?thesis | |
| 66519 | 6066 | unfolding o_def integral_component_eq[OF g] . | 
| 53638 | 6067 | qed | 
| 6068 | ||
| 6069 | lemma has_integral_norm_bound_integral_component: | |
| 56188 | 6070 | fixes f :: "'n::euclidean_space \<Rightarrow> 'a::banach" | 
| 6071 | fixes g :: "'n \<Rightarrow> 'b::euclidean_space" | |
| 66519 | 6072 | assumes f: "(f has_integral i) S" | 
| 6073 | and g: "(g has_integral j) S" | |
| 6074 | and "\<And>x. x \<in> S \<Longrightarrow> norm (f x) \<le> (g x)\<bullet>k" | |
| 53638 | 6075 | shows "norm i \<le> j\<bullet>k" | 
| 66519 | 6076 | using integral_norm_bound_integral_component[of f S g k] | 
| 6077 | unfolding integral_unique[OF f] integral_unique[OF g] | |
| 53638 | 6078 | using assms | 
| 6079 | by auto | |
| 6080 | ||
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6081 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6082 | lemma uniformly_convergent_improper_integral: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6083 |   fixes f :: "'b \<Rightarrow> real \<Rightarrow> 'a :: {banach}"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6084 |   assumes deriv: "\<And>x. x \<ge> a \<Longrightarrow> (G has_field_derivative g x) (at x within {a..})"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6085 |   assumes integrable: "\<And>a' b x. x \<in> A \<Longrightarrow> a' \<ge> a \<Longrightarrow> b \<ge> a' \<Longrightarrow> f x integrable_on {a'..b}"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6086 | assumes G: "convergent G" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6087 | assumes le: "\<And>y x. y \<in> A \<Longrightarrow> x \<ge> a \<Longrightarrow> norm (f y x) \<le> g x" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6088 |   shows   "uniformly_convergent_on A (\<lambda>b x. integral {a..b} (f x))"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6089 | proof (intro Cauchy_uniformly_convergent uniformly_Cauchy_onI', goal_cases) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6090 | case (1 \<epsilon>) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6091 | from G have "Cauchy G" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6092 | by (auto intro!: convergent_Cauchy) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6093 | with 1 obtain M where M: "dist (G (real m)) (G (real n)) < \<epsilon>" if "m \<ge> M" "n \<ge> M" for m n | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6094 | by (force simp: Cauchy_def) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6095 | define M' where "M' = max (nat \<lceil>a\<rceil>) M" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6096 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6097 | show ?case | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6098 | proof (rule exI[of _ M'], safe, goal_cases) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6099 | case (1 x m n) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6100 | have M': "M' \<ge> a" "M' \<ge> M" unfolding M'_def by linarith+ | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6101 |     have int_g: "(g has_integral (G (real n) - G (real m))) {real m..real n}"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6102 | using 1 M' by (intro fundamental_theorem_of_calculus) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6103 | (auto simp: has_field_derivative_iff_has_vector_derivative [symmetric] | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6104 | intro!: DERIV_subset[OF deriv]) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6105 |     have int_f: "f x integrable_on {a'..real n}" if "a' \<ge> a" for a'
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6106 | using that 1 by (cases "a' \<le> real n") (auto intro: integrable) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6107 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6108 |     have "dist (integral {a..real m} (f x)) (integral {a..real n} (f x)) =
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6109 |             norm (integral {a..real n} (f x) - integral {a..real m} (f x))"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6110 | by (simp add: dist_norm norm_minus_commute) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6111 |     also have "integral {a..real m} (f x) + integral {real m..real n} (f x) = 
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6112 |                  integral {a..real n} (f x)"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6113 | using M' and 1 by (intro integral_combine int_f) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6114 |     hence "integral {a..real n} (f x) - integral {a..real m} (f x) = 
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6115 |              integral {real m..real n} (f x)"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6116 | by (simp add: algebra_simps) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6117 |     also have "norm \<dots> \<le> integral {real m..real n} g"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6118 | using le 1 M' int_f int_g by (intro integral_norm_bound_integral) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6119 |     also from int_g have "integral {real m..real n} g = G (real n) - G (real m)"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6120 | by (simp add: has_integral_iff) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6121 | also have "\<dots> \<le> dist (G m) (G n)" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6122 | by (simp add: dist_norm) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6123 | also from 1 and M' have "\<dots> < \<epsilon>" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6124 | by (intro M) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6125 | finally show ?case . | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6126 | qed | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6127 | qed | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6128 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6129 | lemma uniformly_convergent_improper_integral': | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6130 |   fixes f :: "'b \<Rightarrow> real \<Rightarrow> 'a :: {banach, real_normed_algebra}"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6131 |   assumes deriv: "\<And>x. x \<ge> a \<Longrightarrow> (G has_field_derivative g x) (at x within {a..})"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6132 |   assumes integrable: "\<And>a' b x. x \<in> A \<Longrightarrow> a' \<ge> a \<Longrightarrow> b \<ge> a' \<Longrightarrow> f x integrable_on {a'..b}"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6133 | assumes G: "convergent G" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6134 | assumes le: "eventually (\<lambda>x. \<forall>y\<in>A. norm (f y x) \<le> g x) at_top" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6135 |   shows   "uniformly_convergent_on A (\<lambda>b x. integral {a..b} (f x))"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6136 | proof - | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6137 | from le obtain a'' where le: "\<And>y x. y \<in> A \<Longrightarrow> x \<ge> a'' \<Longrightarrow> norm (f y x) \<le> g x" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6138 | by (auto simp: eventually_at_top_linorder) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6139 | define a' where "a' = max a a''" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6140 | |
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6141 |   have "uniformly_convergent_on A (\<lambda>b x. integral {a'..real b} (f x))"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6142 | proof (rule uniformly_convergent_improper_integral) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6143 | fix t assume t: "t \<ge> a'" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6144 |     hence "(G has_field_derivative g t) (at t within {a..})"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6145 | by (intro deriv) (auto simp: a'_def) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6146 |     moreover have "{a'..} \<subseteq> {a..}" unfolding a'_def by auto
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6147 |     ultimately show "(G has_field_derivative g t) (at t within {a'..})"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6148 | by (rule DERIV_subset) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6149 | qed (insert le, auto simp: a'_def intro: integrable G) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6150 |   hence "uniformly_convergent_on A (\<lambda>b x. integral {a..a'} (f x) + integral {a'..real b} (f x))" 
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6151 | (is ?P) by (intro uniformly_convergent_add) auto | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6152 |   also have "eventually (\<lambda>x. \<forall>y\<in>A. integral {a..a'} (f y) + integral {a'..x} (f y) =
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6153 |                    integral {a..x} (f y)) sequentially"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6154 | by (intro eventually_mono [OF eventually_ge_at_top[of "nat \<lceil>a'\<rceil>"]] ballI integral_combine) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6155 | (auto simp: a'_def intro: integrable) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6156 | hence "?P \<longleftrightarrow> ?thesis" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6157 | by (intro uniformly_convergent_cong) simp_all | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6158 | finally show ?thesis . | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6159 | qed | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6160 | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6161 | subsection \<open>differentiation under the integral sign\<close> | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6162 | |
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6163 | lemma integral_continuous_on_param: | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6164 | fixes f::"'a::topological_space \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6165 | assumes cont_fx: "continuous_on (U \<times> cbox a b) (\<lambda>(x, t). f x t)" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6166 | shows "continuous_on U (\<lambda>x. integral (cbox a b) (f x))" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6167 | proof cases | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6168 | assume "content (cbox a b) \<noteq> 0" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6169 |   then have ne: "cbox a b \<noteq> {}" by auto
 | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6170 | |
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6171 | note [continuous_intros] = | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6172 | continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x, | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6173 | unfolded split_beta fst_conv snd_conv] | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6174 | |
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6175 | show ?thesis | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6176 | unfolding continuous_on_def | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6177 | proof (safe intro!: tendstoI) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6178 | fix e'::real and x | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6179 | assume "e' > 0" | 
| 63040 | 6180 | define e where "e = e' / (content (cbox a b) + 1)" | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6181 | have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6182 | assume "x \<in> U" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6183 | from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x \<in> U\<close> \<open>0 < e\<close>] | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6184 | obtain X0 where X0: "x \<in> X0" "open X0" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6185 | and fx_bound: "\<And>y t. y \<in> X0 \<inter> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> norm (f y t - f x t) \<le> e" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6186 | unfolding split_beta fst_conv snd_conv dist_norm | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6187 | by metis | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6188 | have "\<forall>\<^sub>F y in at x within U. y \<in> X0 \<inter> U" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6189 | using X0(1) X0(2) eventually_at_topological by auto | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6190 | then show "\<forall>\<^sub>F y in at x within U. dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6191 | proof eventually_elim | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6192 | case (elim y) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6193 | have "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) = | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6194 | norm (integral (cbox a b) (\<lambda>t. f y t - f x t))" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6195 | using elim \<open>x \<in> U\<close> | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6196 | unfolding dist_norm | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6197 | by (subst integral_diff) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6198 | (auto intro!: integrable_continuous continuous_intros) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6199 | also have "\<dots> \<le> e * content (cbox a b)" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6200 | using elim \<open>x \<in> U\<close> | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6201 | by (intro integrable_bound) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6202 | (auto intro!: fx_bound \<open>x \<in> U \<close> less_imp_le[OF \<open>0 < e\<close>] | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6203 | integrable_continuous continuous_intros) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6204 | also have "\<dots> < e'" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6205 | using \<open>0 < e'\<close> \<open>e > 0\<close> | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6206 | by (auto simp: e_def divide_simps) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6207 | finally show "dist (integral (cbox a b) (f y)) (integral (cbox a b) (f x)) < e'" . | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6208 | qed | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6209 | qed | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6210 | qed (auto intro!: continuous_on_const) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6211 | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6212 | lemma leibniz_rule: | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6213 | fixes f::"'a::banach \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6214 | assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6215 | ((\<lambda>x. f x t) has_derivative blinfun_apply (fx x t)) (at x within U)" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6216 | assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> f x integrable_on cbox a b" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6217 | assumes cont_fx: "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6218 | assumes [intro]: "x0 \<in> U" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6219 | assumes "convex U" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6220 | shows | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6221 | "((\<lambda>x. integral (cbox a b) (f x)) has_derivative integral (cbox a b) (fx x0)) (at x0 within U)" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6222 | (is "(?F has_derivative ?dF) _") | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6223 | proof cases | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6224 | assume "content (cbox a b) \<noteq> 0" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6225 |   then have ne: "cbox a b \<noteq> {}" by auto
 | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6226 | note [continuous_intros] = | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6227 | continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x, | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6228 | unfolded split_beta fst_conv snd_conv] | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6229 | show ?thesis | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6230 | proof (intro has_derivativeI bounded_linear_scaleR_left tendstoI, fold norm_conv_dist) | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6231 | have cont_f1: "\<And>t. t \<in> cbox a b \<Longrightarrow> continuous_on U (\<lambda>x. f x t)" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6232 | by (auto simp: continuous_on_eq_continuous_within intro!: has_derivative_continuous fx) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6233 | note [continuous_intros] = continuous_on_compose2[OF cont_f1] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6234 | fix e'::real | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6235 | assume "e' > 0" | 
| 63040 | 6236 | define e where "e = e' / (content (cbox a b) + 1)" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6237 | have "e > 0" using \<open>e' > 0\<close> by (auto simp: e_def intro!: divide_pos_pos add_nonneg_pos) | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6238 | from continuous_on_prod_compactE[OF cont_fx compact_cbox \<open>x0 \<in> U\<close> \<open>e > 0\<close>] | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6239 | obtain X0 where X0: "x0 \<in> X0" "open X0" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6240 | and fx_bound: "\<And>x t. x \<in> X0 \<inter> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> norm (fx x t - fx x0 t) \<le> e" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6241 | unfolding split_beta fst_conv snd_conv | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6242 | by (metis dist_norm) | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6243 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6244 | note eventually_closed_segment[OF \<open>open X0\<close> \<open>x0 \<in> X0\<close>, of U] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6245 | moreover | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6246 | have "\<forall>\<^sub>F x in at x0 within U. x \<in> X0" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6247 | using \<open>open X0\<close> \<open>x0 \<in> X0\<close> eventually_at_topological by blast | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6248 | moreover have "\<forall>\<^sub>F x in at x0 within U. x \<noteq> x0" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6249 | by (auto simp: eventually_at_filter) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6250 | moreover have "\<forall>\<^sub>F x in at x0 within U. x \<in> U" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6251 | by (auto simp: eventually_at_filter) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6252 | ultimately | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6253 | show "\<forall>\<^sub>F x in at x0 within U. norm ((?F x - ?F x0 - ?dF (x - x0)) /\<^sub>R norm (x - x0)) < e'" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6254 | proof eventually_elim | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6255 | case (elim x) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6256 | from elim have "0 < norm (x - x0)" by simp | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6257 | have "closed_segment x0 x \<subseteq> U" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6258 | by (rule \<open>convex U\<close>[unfolded convex_contains_segment, rule_format, OF \<open>x0 \<in> U\<close> \<open>x \<in> U\<close>]) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6259 | from elim have [intro]: "x \<in> U" by auto | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6260 | have "?F x - ?F x0 - ?dF (x - x0) = | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6261 | integral (cbox a b) (\<lambda>y. f x y - f x0 y - fx x0 y (x - x0))" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6262 | (is "_ = ?id") | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6263 | using \<open>x \<noteq> x0\<close> | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6264 | by (subst blinfun_apply_integral integral_diff, | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6265 | auto intro!: integrable_diff integrable_f2 continuous_intros | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6266 | intro: integrable_continuous)+ | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6267 | also | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6268 |       {
 | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6269 | fix t assume t: "t \<in> (cbox a b)" | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6270 |         have seg: "\<And>t. t \<in> {0..1} \<Longrightarrow> x0 + t *\<^sub>R (x - x0) \<in> X0 \<inter> U"
 | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6271 | using \<open>closed_segment x0 x \<subseteq> U\<close> | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6272 | \<open>closed_segment x0 x \<subseteq> X0\<close> | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6273 | by (force simp: closed_segment_def algebra_simps) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6274 | from t have deriv: | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6275 | "((\<lambda>x. f x t) has_derivative (fx y t)) (at y within X0 \<inter> U)" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6276 | if "y \<in> X0 \<inter> U" for y | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6277 | unfolding has_vector_derivative_def[symmetric] | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6278 | using that \<open>x \<in> X0\<close> | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6279 | by (intro has_derivative_within_subset[OF fx]) auto | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6280 | have "\<forall>x \<in> X0 \<inter> U. onorm (blinfun_apply (fx x t) - (fx x0 t)) \<le> e" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6281 | using fx_bound t | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6282 | by (auto simp add: norm_blinfun_def fun_diff_def blinfun.bilinear_simps[symmetric]) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6283 | from differentiable_bound_linearization[OF seg deriv this] X0 | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6284 | have "norm (f x t - f x0 t - fx x0 t (x - x0)) \<le> e * norm (x - x0)" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6285 | by (auto simp add: ac_simps) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6286 | } | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6287 | then have "norm ?id \<le> integral (cbox a b) (\<lambda>_. e * norm (x - x0))" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6288 | by (intro integral_norm_bound_integral) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6289 | (auto intro!: continuous_intros integrable_diff integrable_f2 | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6290 | intro: integrable_continuous) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6291 | also have "\<dots> = content (cbox a b) * e * norm (x - x0)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6292 | by simp | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6293 | also have "\<dots> < e' * norm (x - x0)" | 
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 6294 | using \<open>e' > 0\<close> | 
| 66154 
bc5e6461f759
Tidying up integration theory and some new theorems
 paulson <lp15@cam.ac.uk> parents: 
66112diff
changeset | 6295 | apply (intro mult_strict_right_mono[OF _ \<open>0 < norm (x - x0)\<close>]) | 
| 66703 | 6296 | apply (auto simp: divide_simps e_def) | 
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 6297 | by (metis \<open>0 < e\<close> e_def order.asym zero_less_divide_iff) | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6298 | finally have "norm (?F x - ?F x0 - ?dF (x - x0)) < e' * norm (x - x0)" . | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6299 | then show ?case | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6300 | by (auto simp: divide_simps) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6301 | qed | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6302 | qed (rule blinfun.bounded_linear_right) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6303 | qed (auto intro!: derivative_eq_intros simp: blinfun.bilinear_simps) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6304 | |
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 6305 | lemma has_vector_derivative_eq_has_derivative_blinfun: | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6306 | "(f has_vector_derivative f') (at x within U) \<longleftrightarrow> | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6307 | (f has_derivative blinfun_scaleR_left f') (at x within U)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6308 | by (simp add: has_vector_derivative_def) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6309 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6310 | lemma leibniz_rule_vector_derivative: | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6311 | fixes f::"real \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6312 | assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6313 | ((\<lambda>x. f x t) has_vector_derivative (fx x t)) (at x within U)" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6314 | assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6315 | assumes cont_fx: "continuous_on (U \<times> cbox a b) (\<lambda>(x, t). fx x t)" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6316 | assumes U: "x0 \<in> U" "convex U" | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6317 | shows "((\<lambda>x. integral (cbox a b) (f x)) has_vector_derivative integral (cbox a b) (fx x0)) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6318 | (at x0 within U)" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6319 | proof - | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6320 | note [continuous_intros] = | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6321 | continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x, | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6322 | unfolded split_beta fst_conv snd_conv] | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6323 | have *: "blinfun_scaleR_left (integral (cbox a b) (fx x0)) = | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6324 | integral (cbox a b) (\<lambda>t. blinfun_scaleR_left (fx x0 t))" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6325 | by (subst integral_linear[symmetric]) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6326 | (auto simp: has_vector_derivative_def o_def | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6327 | intro!: integrable_continuous U continuous_intros bounded_linear_intros) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6328 | show ?thesis | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6329 | unfolding has_vector_derivative_eq_has_derivative_blinfun | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6330 | apply (rule has_derivative_eq_rhs) | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6331 | apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\<lambda>x t. blinfun_scaleR_left (fx x t)"]) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6332 | using fx cont_fx | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6333 | apply (auto simp: has_vector_derivative_def * split_beta intro!: continuous_intros) | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6334 | done | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6335 | qed | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6336 | |
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 6337 | lemma has_field_derivative_eq_has_derivative_blinfun: | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6338 | "(f has_field_derivative f') (at x within U) \<longleftrightarrow> (f has_derivative blinfun_mult_right f') (at x within U)" | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6339 | by (simp add: has_field_derivative_def) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6340 | |
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6341 | lemma leibniz_rule_field_derivative: | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6342 |   fixes f::"'a::{real_normed_field, banach} \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'a"
 | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6343 | assumes fx: "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> ((\<lambda>x. f x t) has_field_derivative fx x t) (at x within U)" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6344 | assumes integrable_f2: "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6345 | assumes cont_fx: "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6346 | assumes U: "x0 \<in> U" "convex U" | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6347 | shows "((\<lambda>x. integral (cbox a b) (f x)) has_field_derivative integral (cbox a b) (fx x0)) (at x0 within U)" | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6348 | proof - | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6349 | note [continuous_intros] = | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6350 | continuous_on_compose2[OF cont_fx, where f="\<lambda>y. Pair x y" for x, | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6351 | unfolded split_beta fst_conv snd_conv] | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6352 | have *: "blinfun_mult_right (integral (cbox a b) (fx x0)) = | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6353 | integral (cbox a b) (\<lambda>t. blinfun_mult_right (fx x0 t))" | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6354 | by (subst integral_linear[symmetric]) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6355 | (auto simp: has_vector_derivative_def o_def | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6356 | intro!: integrable_continuous U continuous_intros bounded_linear_intros) | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6357 | show ?thesis | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6358 | unfolding has_field_derivative_eq_has_derivative_blinfun | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6359 | apply (rule has_derivative_eq_rhs) | 
| 62182 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6360 | apply (rule leibniz_rule[OF _ integrable_f2 _ U, where fx="\<lambda>x t. blinfun_mult_right (fx x t)"]) | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6361 | using fx cont_fx | 
| 
9ca00b65d36c
continuity of parameterized integral; easier-to-apply formulation of rules
 immler parents: 
61973diff
changeset | 6362 | apply (auto simp: has_field_derivative_def * split_beta intro!: continuous_intros) | 
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6363 | done | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6364 | qed | 
| 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6365 | |
| 67371 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6366 | lemma leibniz_rule_field_differentiable: | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6367 |   fixes f::"'a::{real_normed_field, banach} \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'a"
 | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6368 | assumes "\<And>x t. x \<in> U \<Longrightarrow> t \<in> cbox a b \<Longrightarrow> ((\<lambda>x. f x t) has_field_derivative fx x t) (at x within U)" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6369 | assumes "\<And>x. x \<in> U \<Longrightarrow> (f x) integrable_on cbox a b" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6370 | assumes "continuous_on (U \<times> (cbox a b)) (\<lambda>(x, t). fx x t)" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6371 | assumes "x0 \<in> U" "convex U" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6372 | shows "(\<lambda>x. integral (cbox a b) (f x)) field_differentiable at x0 within U" | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6373 | using leibniz_rule_field_derivative[OF assms] by (auto simp: field_differentiable_def) | 
| 
2d9cf74943e1
moved in some material from Euler-MacLaurin
 paulson <lp15@cam.ac.uk> parents: 
66884diff
changeset | 6374 | |
| 61915 
e9812a95d108
theory for type of bounded linear functions; differentiation under the integral sign
 immler parents: 
61824diff
changeset | 6375 | |
| 61243 | 6376 | subsection \<open>Exchange uniform limit and integral\<close> | 
| 6377 | ||
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6378 | lemma uniform_limit_integral_cbox: | 
| 61243 | 6379 | fixes f::"'a \<Rightarrow> 'b::euclidean_space \<Rightarrow> 'c::banach" | 
| 6380 | assumes u: "uniform_limit (cbox a b) f g F" | |
| 6381 | assumes c: "\<And>n. continuous_on (cbox a b) (f n)" | |
| 6382 | assumes [simp]: "F \<noteq> bot" | |
| 6383 | obtains I J where | |
| 6384 | "\<And>n. (f n has_integral I n) (cbox a b)" | |
| 6385 | "(g has_integral J) (cbox a b)" | |
| 61973 | 6386 | "(I \<longlongrightarrow> J) F" | 
| 61243 | 6387 | proof - | 
| 6388 | have fi[simp]: "f n integrable_on (cbox a b)" for n | |
| 6389 | by (auto intro!: integrable_continuous assms) | |
| 6390 | then obtain I where I: "\<And>n. (f n has_integral I n) (cbox a b)" | |
| 6391 | by atomize_elim (auto simp: integrable_on_def intro!: choice) | |
| 6392 | ||
| 6393 | moreover | |
| 6394 | have gi[simp]: "g integrable_on (cbox a b)" | |
| 6395 | by (auto intro!: integrable_continuous uniform_limit_theorem[OF _ u] eventuallyI c) | |
| 6396 | then obtain J where J: "(g has_integral J) (cbox a b)" | |
| 6397 | by blast | |
| 6398 | ||
| 6399 | moreover | |
| 61973 | 6400 | have "(I \<longlongrightarrow> J) F" | 
| 61243 | 6401 | proof cases | 
| 6402 | assume "content (cbox a b) = 0" | |
| 6403 | hence "I = (\<lambda>_. 0)" "J = 0" | |
| 6404 | by (auto intro!: has_integral_unique I J) | |
| 6405 | thus ?thesis by simp | |
| 6406 | next | |
| 6407 | assume content_nonzero: "content (cbox a b) \<noteq> 0" | |
| 6408 | show ?thesis | |
| 6409 | proof (rule tendstoI) | |
| 6410 | fix e::real | |
| 6411 | assume "e > 0" | |
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6412 | define e' where "e' = e/2" | 
| 61243 | 6413 | with \<open>e > 0\<close> have "e' > 0" by simp | 
| 6414 | then have "\<forall>\<^sub>F n in F. \<forall>x\<in>cbox a b. norm (f n x - g x) < e' / content (cbox a b)" | |
| 66089 
def95e0bc529
Some new material. SIMPRULE STATUS for sum/prod.delta rules!
 paulson <lp15@cam.ac.uk> parents: 
65680diff
changeset | 6415 | using u content_nonzero by (auto simp: uniform_limit_iff dist_norm zero_less_measure_iff) | 
| 61243 | 6416 | then show "\<forall>\<^sub>F n in F. dist (I n) J < e" | 
| 6417 | proof eventually_elim | |
| 6418 | case (elim n) | |
| 6419 | have "I n = integral (cbox a b) (f n)" | |
| 6420 | "J = integral (cbox a b) g" | |
| 6421 | using I[of n] J by (simp_all add: integral_unique) | |
| 6422 | then have "dist (I n) J = norm (integral (cbox a b) (\<lambda>x. f n x - g x))" | |
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 6423 | by (simp add: integral_diff dist_norm) | 
| 61243 | 6424 | also have "\<dots> \<le> integral (cbox a b) (\<lambda>x. (e' / content (cbox a b)))" | 
| 6425 | using elim | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63940diff
changeset | 6426 | by (intro integral_norm_bound_integral) (auto intro!: integrable_diff) | 
| 61243 | 6427 | also have "\<dots> < e" | 
| 6428 | using \<open>0 < e\<close> | |
| 6429 | by (simp add: e'_def) | |
| 6430 | finally show ?case . | |
| 6431 | qed | |
| 6432 | qed | |
| 6433 | qed | |
| 6434 | ultimately show ?thesis .. | |
| 6435 | qed | |
| 6436 | ||
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6437 | lemma uniform_limit_integral: | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6438 | fixes f::"'a \<Rightarrow> 'b::ordered_euclidean_space \<Rightarrow> 'c::banach" | 
| 66402 | 6439 |   assumes u: "uniform_limit {a..b} f g F"
 | 
| 6440 |   assumes c: "\<And>n. continuous_on {a..b} (f n)"
 | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6441 | assumes [simp]: "F \<noteq> bot" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6442 | obtains I J where | 
| 66402 | 6443 |     "\<And>n. (f n has_integral I n) {a..b}"
 | 
| 6444 |     "(g has_integral J) {a..b}"
 | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6445 | "(I \<longlongrightarrow> J) F" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6446 | by (metis interval_cbox assms uniform_limit_integral_cbox) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6447 | |
| 61243 | 6448 | |
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6449 | subsection \<open>Integration by parts\<close> | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6450 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6451 | lemma integration_by_parts_interior_strong: | 
| 64272 | 6452 | fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach" | 
| 6453 | assumes bilinear: "bounded_bilinear (prod)" | |
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6454 | assumes s: "finite s" and le: "a \<le> b" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6455 |   assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6456 |   assumes deriv: "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6457 |                  "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
 | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 6458 | assumes int: "((\<lambda>x. prod (f x) (g' x)) has_integral | 
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6459 |                   (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6460 |   shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6461 | proof - | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6462 | interpret bounded_bilinear prod by fact | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6463 | have "((\<lambda>x. prod (f x) (g' x) + prod (f' x) (g x)) has_integral | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6464 |           (prod (f b) (g b) - prod (f a) (g a))) {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6465 | using deriv by (intro fundamental_theorem_of_calculus_interior_strong[OF s le]) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6466 | (auto intro!: continuous_intros continuous_on has_vector_derivative) | 
| 66112 
0e640e04fc56
New theorems; stronger theorems; tidier theorems. Also some renaming
 paulson <lp15@cam.ac.uk> parents: 
66089diff
changeset | 6467 | from has_integral_diff[OF this int] show ?thesis by (simp add: algebra_simps) | 
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6468 | qed | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6469 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6470 | lemma integration_by_parts_interior: | 
| 64272 | 6471 | fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach" | 
| 6472 | assumes "bounded_bilinear (prod)" "a \<le> b" | |
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6473 |           "continuous_on {a..b} f" "continuous_on {a..b} g"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6474 |   assumes "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6475 |           "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6476 |   assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6477 |   shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6478 |   by (rule integration_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6479 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6480 | lemma integration_by_parts: | 
| 64272 | 6481 | fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach" | 
| 6482 | assumes "bounded_bilinear (prod)" "a \<le> b" | |
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6483 |           "continuous_on {a..b} f" "continuous_on {a..b} g"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6484 |   assumes "\<And>x. x\<in>{a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6485 |           "\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6486 |   assumes "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - y)) {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6487 |   shows   "((\<lambda>x. prod (f' x) (g x)) has_integral y) {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6488 | by (rule integration_by_parts_interior[of _ _ _ f g f' g']) (insert assms, simp_all) | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6489 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6490 | lemma integrable_by_parts_interior_strong: | 
| 64272 | 6491 | fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach" | 
| 6492 | assumes bilinear: "bounded_bilinear (prod)" | |
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6493 | assumes s: "finite s" and le: "a \<le> b" | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6494 |   assumes cont [continuous_intros]: "continuous_on {a..b} f" "continuous_on {a..b} g"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6495 |   assumes deriv: "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6496 |                  "\<And>x. x\<in>{a<..<b} - s \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6497 |   assumes int: "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6498 |   shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6499 | proof - | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6500 |   from int obtain I where "((\<lambda>x. prod (f x) (g' x)) has_integral I) {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6501 | unfolding integrable_on_def by blast | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6502 | hence "((\<lambda>x. prod (f x) (g' x)) has_integral (prod (f b) (g b) - prod (f a) (g a) - | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6503 |            (prod (f b) (g b) - prod (f a) (g a) - I))) {a..b}" by simp
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6504 | from integration_by_parts_interior_strong[OF assms(1-7) this] | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6505 | show ?thesis unfolding integrable_on_def by blast | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6506 | qed | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6507 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6508 | lemma integrable_by_parts_interior: | 
| 64272 | 6509 | fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach" | 
| 6510 | assumes "bounded_bilinear (prod)" "a \<le> b" | |
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6511 |           "continuous_on {a..b} f" "continuous_on {a..b} g"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6512 |   assumes "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6513 |           "\<And>x. x\<in>{a<..<b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6514 |   assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6515 |   shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6516 |   by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6517 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6518 | lemma integrable_by_parts: | 
| 64272 | 6519 | fixes prod :: "_ \<Rightarrow> _ \<Rightarrow> 'b :: banach" | 
| 6520 | assumes "bounded_bilinear (prod)" "a \<le> b" | |
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6521 |           "continuous_on {a..b} f" "continuous_on {a..b} g"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6522 |   assumes "\<And>x. x\<in>{a..b} \<Longrightarrow> (f has_vector_derivative f' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6523 |           "\<And>x. x\<in>{a..b} \<Longrightarrow> (g has_vector_derivative g' x) (at x)"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6524 |   assumes "(\<lambda>x. prod (f x) (g' x)) integrable_on {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6525 |   shows   "(\<lambda>x. prod (f' x) (g x)) integrable_on {a..b}"
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6526 |   by (rule integrable_by_parts_interior_strong[of _ "{}" _ _ f g f' g']) (insert assms, simp_all)
 | 
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6527 | |
| 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6528 | |
| 63299 | 6529 | subsection \<open>Integration by substitution\<close> | 
| 6530 | ||
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6531 | lemma has_integral_substitution_general: | 
| 63299 | 6532 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real" | 
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6533 | assumes s: "finite s" and le: "a \<le> b" | 
| 63299 | 6534 |       and subset: "g ` {a..b} \<subseteq> {c..d}"
 | 
| 6535 |       and f [continuous_intros]: "continuous_on {c..d} f"
 | |
| 6536 |       and g [continuous_intros]: "continuous_on {a..b} g"
 | |
| 6537 | and deriv [derivative_intros]: | |
| 6538 |               "\<And>x. x \<in> {a..b} - s \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
 | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6539 |     shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f - integral {g b..g a} f)) {a..b}"
 | 
| 63299 | 6540 | proof - | 
| 6541 |   let ?F = "\<lambda>x. integral {c..g x} f"
 | |
| 6542 |   have cont_int: "continuous_on {a..b} ?F"
 | |
| 66192 
e5b84854baa4
A few renamings and several tidied-up proofs
 paulson <lp15@cam.ac.uk> parents: 
66164diff
changeset | 6543 | by (rule continuous_on_compose2[OF _ g subset] indefinite_integral_continuous_1 | 
| 63299 | 6544 | f integrable_continuous_real)+ | 
| 6545 |   have deriv: "(((\<lambda>x. integral {c..x} f) \<circ> g) has_vector_derivative g' x *\<^sub>R f (g x))
 | |
| 6546 |                  (at x within {a..b})" if "x \<in> {a..b} - s" for x
 | |
| 66703 | 6547 | proof (rule has_vector_derivative_eq_rhs [OF vector_diff_chain_within refl]) | 
| 6548 |     show "(g has_vector_derivative g' x) (at x within {a..b})"
 | |
| 6549 | using deriv has_field_derivative_iff_has_vector_derivative that by blast | |
| 6550 |     show "((\<lambda>x. integral {c..x} f) has_vector_derivative f (g x)) 
 | |
| 6551 |           (at (g x) within g ` {a..b})"
 | |
| 6552 | using that le subset | |
| 6553 | by (blast intro: has_vector_derivative_within_subset integral_has_vector_derivative f) | |
| 6554 | qed | |
| 63299 | 6555 | have deriv: "(?F has_vector_derivative g' x *\<^sub>R f (g x)) | 
| 6556 |                   (at x)" if "x \<in> {a..b} - (s \<union> {a,b})" for x
 | |
| 66793 
deabce3ccf1f
new material about connectedness, etc.
 paulson <lp15@cam.ac.uk> parents: 
66709diff
changeset | 6557 | using deriv[of x] that by (simp add: at_within_Icc_at o_def) | 
| 63299 | 6558 |   have "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (?F b - ?F a)) {a..b}"
 | 
| 6559 | using le cont_int s deriv cont_int | |
| 6560 |     by (intro fundamental_theorem_of_calculus_interior_strong[of "s \<union> {a,b}"]) simp_all
 | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6561 | also | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6562 |   from subset have "g x \<in> {c..d}" if "x \<in> {a..b}" for x using that by blast
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6563 | from this[of a] this[of b] le have cd: "c \<le> g a" "g b \<le> d" "c \<le> g b" "g a \<le> d" by auto | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6564 |   have "integral {c..g b} f - integral {c..g a} f = integral {g a..g b} f - integral {g b..g a} f"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6565 | proof cases | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6566 | assume "g a \<le> g b" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6567 | note le = le this | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6568 |     from cd have "integral {c..g a} f + integral {g a..g b} f = integral {c..g b} f"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6569 | by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6570 | with le show ?thesis | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6571 | by (cases "g a = g b") (simp_all add: algebra_simps) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6572 | next | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6573 | assume less: "\<not>g a \<le> g b" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6574 | then have "g a \<ge> g b" by simp | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6575 | note le = le this | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6576 |     from cd have "integral {c..g b} f + integral {g b..g a} f = integral {c..g a} f"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6577 | by (intro integral_combine integrable_continuous_real continuous_on_subset[OF f] le) simp_all | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6578 | with less show ?thesis | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6579 | by (simp_all add: algebra_simps) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6580 | qed | 
| 63299 | 6581 | finally show ?thesis . | 
| 6582 | qed | |
| 6583 | ||
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6584 | lemma has_integral_substitution_strong: | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6585 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6586 | assumes s: "finite s" and le: "a \<le> b" "g a \<le> g b" | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6587 |     and subset: "g ` {a..b} \<subseteq> {c..d}"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6588 |     and f [continuous_intros]: "continuous_on {c..d} f"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6589 |     and g [continuous_intros]: "continuous_on {a..b} g"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6590 | and deriv [derivative_intros]: | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6591 |     "\<And>x. x \<in> {a..b} - s \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6592 |   shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}"
 | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6593 | using has_integral_substitution_general[OF s le(1) subset f g deriv] le(2) | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6594 | by (cases "g a = g b") auto | 
| 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6595 | |
| 63299 | 6596 | lemma has_integral_substitution: | 
| 6597 | fixes f :: "real \<Rightarrow> 'a::euclidean_space" and g :: "real \<Rightarrow> real" | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6598 |   assumes "a \<le> b" "g a \<le> g b" "g ` {a..b} \<subseteq> {c..d}"
 | 
| 63299 | 6599 |       and "continuous_on {c..d} f"
 | 
| 6600 |       and "\<And>x. x \<in> {a..b} \<Longrightarrow> (g has_field_derivative g' x) (at x within {a..b})"
 | |
| 6601 |     shows "((\<lambda>x. g' x *\<^sub>R f (g x)) has_integral (integral {g a..g b} f)) {a..b}"
 | |
| 65204 
d23eded35a33
modernized construction of type bcontfun; base explicit theorems on Uniform_Limit.thy; added some lemmas
 immler parents: 
65036diff
changeset | 6602 |   by (intro has_integral_substitution_strong[of "{}" a b g c d] assms)
 | 
| 63299 | 6603 | (auto intro: DERIV_continuous_on assms) | 
| 6604 | ||
| 6605 | ||
| 63295 
52792bb9126e
Facts about HK integration, complex powers, Gamma function
 eberlm parents: 
63170diff
changeset | 6606 | subsection \<open>Compute a double integral using iterated integrals and switching the order of integration\<close> | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6607 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6608 | lemma continuous_on_imp_integrable_on_Pair1: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6609 | fixes f :: "_ \<Rightarrow> 'b::banach" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6610 | assumes con: "continuous_on (cbox (a,c) (b,d)) f" and x: "x \<in> cbox a b" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6611 | shows "(\<lambda>y. f (x, y)) integrable_on (cbox c d)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6612 | proof - | 
| 61736 | 6613 | have "f \<circ> (\<lambda>y. (x, y)) integrable_on (cbox c d)" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6614 | apply (rule integrable_continuous) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6615 | apply (rule continuous_on_compose [OF _ continuous_on_subset [OF con]]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6616 | using x | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6617 | apply (auto intro: continuous_on_Pair continuous_on_const continuous_on_id continuous_on_subset con) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6618 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6619 | then show ?thesis | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6620 | by (simp add: o_def) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6621 | qed | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6622 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6623 | lemma integral_integrable_2dim: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6624 |   fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6625 | assumes "continuous_on (cbox (a,c) (b,d)) f" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6626 | shows "(\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y))) integrable_on cbox a b" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6627 | proof (cases "content(cbox c d) = 0") | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6628 | case True | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6629 | then show ?thesis | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6630 | by (simp add: True integrable_const) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6631 | next | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6632 | case False | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6633 | have uc: "uniformly_continuous_on (cbox (a,c) (b,d)) f" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6634 | by (simp add: assms compact_cbox compact_uniformly_continuous) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6635 |   { fix x::'a and e::real
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6636 | assume x: "x \<in> cbox a b" and e: "0 < e" | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6637 | then have e2_gt: "0 < e/2 / content (cbox c d)" and e2_less: "e/2 / content (cbox c d) * content (cbox c d) < e" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6638 | by (auto simp: False content_lt_nz e) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6639 | then obtain dd | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6640 | where dd: "\<And>x x'. \<lbrakk>x\<in>cbox (a, c) (b, d); x'\<in>cbox (a, c) (b, d); norm (x' - x) < dd\<rbrakk> | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 6641 | \<Longrightarrow> norm (f x' - f x) \<le> e/(2 * content (cbox c d))" "dd>0" | 
| 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 6642 | using uc [unfolded uniformly_continuous_on_def, THEN spec, of "e/(2 * content (cbox c d))"] | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6643 | by (auto simp: dist_norm intro: less_imp_le) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6644 | have "\<exists>delta>0. \<forall>x'\<in>cbox a b. norm (x' - x) < delta \<longrightarrow> norm (integral (cbox c d) (\<lambda>u. f (x', u) - f (x, u))) < e" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6645 | apply (rule_tac x=dd in exI) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6646 | using dd e2_gt assms x | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6647 | apply clarify | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6648 | apply (rule le_less_trans [OF _ e2_less]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6649 | apply (rule integrable_bound) | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 6650 | apply (auto intro: integrable_diff continuous_on_imp_integrable_on_Pair1) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6651 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6652 | } note * = this | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6653 | show ?thesis | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6654 | apply (rule integrable_continuous) | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 6655 | apply (simp add: * continuous_on_iff dist_norm integral_diff [symmetric] continuous_on_imp_integrable_on_Pair1 [OF assms]) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6656 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6657 | qed | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6658 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6659 | lemma integral_split: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6660 |   fixes f :: "'a::euclidean_space \<Rightarrow> 'b::{real_normed_vector,complete_space}"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6661 | assumes f: "f integrable_on (cbox a b)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6662 | and k: "k \<in> Basis" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6663 | shows "integral (cbox a b) f = | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6664 |            integral (cbox a b \<inter> {x. x\<bullet>k \<le> c}) f +
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6665 |            integral (cbox a b \<inter> {x. x\<bullet>k \<ge> c}) f"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6666 | apply (rule integral_unique [OF has_integral_split [where c=c]]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6667 | using k f | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6668 | apply (auto simp: has_integral_integral [symmetric]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6669 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6670 | |
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6671 | lemma integral_swap_operativeI: | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6672 |   fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6673 | assumes f: "continuous_on s f" and e: "0 < e" | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6674 | shows "operative conj True | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6675 | (\<lambda>k. \<forall>a b c d. | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6676 | cbox (a,c) (b,d) \<subseteq> k \<and> cbox (a,c) (b,d) \<subseteq> s | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6677 | \<longrightarrow> norm(integral (cbox (a,c) (b,d)) f - | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6678 | integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f((x,y))))) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6679 | \<le> e * content (cbox (a,c) (b,d)))" | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6680 | proof (standard, auto) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6681 | fix a::'a and c::'b and b::'a and d::'b and u::'a and v::'a and w::'b and z::'b | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 6682 |   assume *: "box (a, c) (b, d) = {}"
 | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6683 | and cb1: "cbox (u, w) (v, z) \<subseteq> cbox (a, c) (b, d)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6684 | and cb2: "cbox (u, w) (v, z) \<subseteq> s" | 
| 63957 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 6685 | then have c0: "content (cbox (a, c) (b, d)) = 0" | 
| 
c3da799b1b45
HOL-Analysis: move gauges and (tagged) divisions to its own theory file
 hoelzl parents: 
63956diff
changeset | 6686 | using * unfolding content_eq_0_interior by simp | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6687 | have c0': "content (cbox (u, w) (v, z)) = 0" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6688 | by (fact content_0_subset [OF c0 cb1]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6689 | show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y)))) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6690 | \<le> e * content (cbox (u,w) (v,z))" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6691 | using content_cbox_pair_eq0_D [OF c0'] | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6692 | by (force simp add: c0') | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6693 | next | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6694 | fix a::'a and c::'b and b::'a and d::'b | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6695 | and M::real and i::'a and j::'b | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6696 | and u::'a and v::'a and w::'b and z::'b | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6697 | assume ij: "(i,j) \<in> Basis" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6698 | and n1: "\<forall>a' b' c' d'. | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6699 | cbox (a',c') (b',d') \<subseteq> cbox (a,c) (b,d) \<and> | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6700 |                 cbox (a',c') (b',d') \<subseteq> {x. x \<bullet> (i,j) \<le> M} \<and> cbox (a',c') (b',d') \<subseteq> s \<longrightarrow>
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6701 | norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f (x,y)))) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6702 | \<le> e * content (cbox (a',c') (b',d'))" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6703 | and n2: "\<forall>a' b' c' d'. | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6704 | cbox (a',c') (b',d') \<subseteq> cbox (a,c) (b,d) \<and> | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6705 |                 cbox (a',c') (b',d') \<subseteq> {x. M \<le> x \<bullet> (i,j)} \<and> cbox (a',c') (b',d') \<subseteq> s \<longrightarrow>
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6706 | norm (integral (cbox (a',c') (b',d')) f - integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f (x,y)))) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6707 | \<le> e * content (cbox (a',c') (b',d'))" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6708 | and subs: "cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)" "cbox (u,w) (v,z) \<subseteq> s" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6709 | have fcont: "continuous_on (cbox (u, w) (v, z)) f" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6710 | using assms(1) continuous_on_subset subs(2) by blast | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6711 | then have fint: "f integrable_on cbox (u, w) (v, z)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6712 | by (metis integrable_continuous) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6713 | consider "i \<in> Basis" "j=0" | "j \<in> Basis" "i=0" using ij | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6714 | by (auto simp: Euclidean_Space.Basis_prod_def) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6715 | then show "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y)))) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6716 | \<le> e * content (cbox (u,w) (v,z))" (is ?normle) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6717 | proof cases | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6718 | case 1 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6719 | then have e: "e * content (cbox (u, w) (v, z)) = | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6720 |                   e * (content (cbox u v \<inter> {x. x \<bullet> i \<le> M}) * content (cbox w z)) +
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6721 |                   e * (content (cbox u v \<inter> {x. M \<le> x \<bullet> i}) * content (cbox w z))"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6722 | by (simp add: content_split [where c=M] content_Pair algebra_simps) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6723 | have *: "integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) = | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6724 |                 integral (cbox u v \<inter> {x. x \<bullet> i \<le> M}) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) +
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6725 |                 integral (cbox u v \<inter> {x. M \<le> x \<bullet> i}) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y)))"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6726 | using 1 f subs integral_integrable_2dim continuous_on_subset | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6727 | by (blast intro: integral_split) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6728 | show ?normle | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6729 | apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6730 | using 1 subs | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6731 | apply (simp_all add: cbox_Pair_eq setcomp_dot1 [of "\<lambda>u. M\<le>u"] setcomp_dot1 [of "\<lambda>u. u\<le>M"] Sigma_Int_Paircomp1) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6732 | apply (simp_all add: interval_split ij) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6733 | apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6734 | apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6735 | apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6736 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6737 | next | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6738 | case 2 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6739 | then have e: "e * content (cbox (u, w) (v, z)) = | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6740 |                   e * (content (cbox u v) * content (cbox w z \<inter> {x. x \<bullet> j \<le> M})) +
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6741 |                   e * (content (cbox u v) * content (cbox w z \<inter> {x. M \<le> x \<bullet> j}))"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6742 | by (simp add: content_split [where c=M] content_Pair algebra_simps) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6743 |     have "(\<lambda>x. integral (cbox w z \<inter> {x. x \<bullet> j \<le> M}) (\<lambda>y. f (x, y))) integrable_on cbox u v"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6744 |                 "(\<lambda>x. integral (cbox w z \<inter> {x. M \<le> x \<bullet> j}) (\<lambda>y. f (x, y))) integrable_on cbox u v"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6745 | using 2 subs | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6746 | apply (simp_all add: interval_split) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6747 | apply (rule_tac [!] integral_integrable_2dim [OF continuous_on_subset [OF f]]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6748 | apply (auto simp: cbox_Pair_eq interval_split [symmetric]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6749 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6750 | with 2 have *: "integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) = | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6751 |                    integral (cbox u v) (\<lambda>x. integral (cbox w z \<inter> {x. x \<bullet> j \<le> M}) (\<lambda>y. f (x, y))) +
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6752 |                    integral (cbox u v) (\<lambda>x. integral (cbox w z \<inter> {x. M \<le> x \<bullet> j}) (\<lambda>y. f (x, y)))"
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6753 | by (simp add: integral_add [symmetric] integral_split [symmetric] | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6754 | continuous_on_imp_integrable_on_Pair1 [OF fcont] cong: integral_cong) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6755 | show ?normle | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6756 | apply (rule norm_diff2 [OF integral_split [where c=M, OF fint ij] * e]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6757 | using 2 subs | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6758 | apply (simp_all add: cbox_Pair_eq setcomp_dot2 [of "\<lambda>u. M\<le>u"] setcomp_dot2 [of "\<lambda>u. u\<le>M"] Sigma_Int_Paircomp2) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6759 | apply (simp_all add: interval_split ij) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6760 | apply (simp_all add: cbox_Pair_eq [symmetric] content_Pair [symmetric]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6761 | apply (force simp add: interval_split [symmetric] intro!: n1 [rule_format]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6762 | apply (force simp add: interval_split [symmetric] intro!: n2 [rule_format]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6763 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6764 | qed | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6765 | qed | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6766 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6767 | lemma integral_Pair_const: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6768 | "integral (cbox (a,c) (b,d)) (\<lambda>x. k) = | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6769 | integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. k))" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6770 | by (simp add: content_Pair) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6771 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6772 | lemma integral_prod_continuous: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6773 |   fixes f :: "('a::euclidean_space * 'b::euclidean_space) \<Rightarrow> 'c::banach"
 | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6774 | assumes "continuous_on (cbox (a, c) (b, d)) f" (is "continuous_on ?CBOX f") | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6775 | shows "integral (cbox (a, c) (b, d)) f = integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x, y)))" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6776 | proof (cases "content ?CBOX = 0") | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6777 | case True | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6778 | then show ?thesis | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6779 | by (auto simp: content_Pair) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6780 | next | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6781 | case False | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6782 | then have cbp: "content ?CBOX > 0" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6783 | using content_lt_nz by blast | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6784 | have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) = 0" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6785 | proof (rule dense_eq0_I, simp) | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6786 | fix e :: real | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6787 | assume "0 < e" | 
| 66505 | 6788 | with \<open>content ?CBOX > 0\<close> have "0 < e/content ?CBOX" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6789 | by simp | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6790 | have f_int_acbd: "f integrable_on ?CBOX" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6791 | by (rule integrable_continuous [OF assms]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6792 |     { fix p
 | 
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6793 | assume p: "p division_of ?CBOX" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6794 | then have "finite p" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6795 | by blast | 
| 66505 | 6796 | define e' where "e' = e/content ?CBOX" | 
| 6797 | with \<open>0 < e\<close> \<open>0 < e/content ?CBOX\<close> | |
| 66492 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6798 | have "0 < e'" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6799 | by simp | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6800 | interpret operative conj True | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6801 | "\<lambda>k. \<forall>a' b' c' d'. | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6802 | cbox (a', c') (b', d') \<subseteq> k \<and> cbox (a', c') (b', d') \<subseteq> ?CBOX | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6803 | \<longrightarrow> norm (integral (cbox (a', c') (b', d')) f - | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6804 | integral (cbox a' b') (\<lambda>x. integral (cbox c' d') (\<lambda>y. f ((x, y))))) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6805 | \<le> e' * content (cbox (a', c') (b', d'))" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6806 | using assms \<open>0 < e'\<close> by (rule integral_swap_operativeI) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6807 | have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x, y)))) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6808 | \<le> e' * content ?CBOX" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6809 | if "\<And>t u v w z. t \<in> p \<Longrightarrow> cbox (u, w) (v, z) \<subseteq> t \<Longrightarrow> cbox (u, w) (v, z) \<subseteq> ?CBOX | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6810 | \<Longrightarrow> norm (integral (cbox (u, w) (v, z)) f - | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6811 | integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y)))) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6812 | \<le> e' * content (cbox (u, w) (v, z))" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6813 | using that division [of p "(a, c)" "(b, d)"] p \<open>finite p\<close> by (auto simp add: comm_monoid_set_F_and) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6814 | with False have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x, y)))) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6815 | \<le> e" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6816 | if "\<And>t u v w z. t \<in> p \<Longrightarrow> cbox (u, w) (v, z) \<subseteq> t \<Longrightarrow> cbox (u, w) (v, z) \<subseteq> ?CBOX | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6817 | \<Longrightarrow> norm (integral (cbox (u, w) (v, z)) f - | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6818 | integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y)))) | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6819 | \<le> e * content (cbox (u, w) (v, z)) / content ?CBOX" | 
| 
d7206afe2d28
dedicated local for "operative" avoids namespace pollution
 haftmann parents: 
66487diff
changeset | 6820 | using that by (simp add: e'_def) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6821 | } note op_acbd = this | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 6822 |     { fix k::real and \<D> and u::'a and v w and z::'b and t1 t2 l
 | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6823 | assume k: "0 < k" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6824 | and nf: "\<And>x y u v. | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6825 | \<lbrakk>x \<in> cbox a b; y \<in> cbox c d; u \<in> cbox a b; v\<in>cbox c d; norm (u-x, v-y) < k\<rbrakk> | 
| 66487 
307c19f24d5c
more on the dreadful monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66437diff
changeset | 6826 | \<Longrightarrow> norm (f(u,v) - f(x,y)) < e/(2 * (content ?CBOX))" | 
| 66495 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 6827 | and p_acbd: "\<D> tagged_division_of cbox (a,c) (b,d)" | 
| 
0b46bd081228
More tidying up of monotone_convergence_interval
 paulson <lp15@cam.ac.uk> parents: 
66487diff
changeset | 6828 | and fine: "(\<lambda>x. ball x k) fine \<D>" "((t1,t2), l) \<in> \<D>" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6829 | and uwvz_sub: "cbox (u,w) (v,z) \<subseteq> l" "cbox (u,w) (v,z) \<subseteq> cbox (a,c) (b,d)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6830 | have t: "t1 \<in> cbox a b" "t2 \<in> cbox c d" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6831 | by (meson fine p_acbd cbox_Pair_iff tag_in_interval)+ | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6832 | have ls: "l \<subseteq> ball (t1,t2) k" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6833 | using fine by (simp add: fine_def Ball_def) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6834 |       { fix x1 x2
 | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6835 | assume xuvwz: "x1 \<in> cbox u v" "x2 \<in> cbox w z" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6836 | then have x: "x1 \<in> cbox a b" "x2 \<in> cbox c d" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6837 | using uwvz_sub by auto | 
| 65036 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 6838 | have "norm (x1 - t1, x2 - t2) = norm (t1 - x1, t2 - x2)" | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 6839 | by (simp add: norm_Pair norm_minus_commute) | 
| 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 6840 | also have "norm (t1 - x1, t2 - x2) < k" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6841 | using xuvwz ls uwvz_sub unfolding ball_def | 
| 65036 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 6842 | by (force simp add: cbox_Pair_eq dist_norm ) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 6843 | finally have "norm (f (x1,x2) - f (t1,t2)) \<le> e/content ?CBOX/2" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6844 | using nf [OF t x] by simp | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6845 | } note nf' = this | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6846 | have f_int_uwvz: "f integrable_on cbox (u,w) (v,z)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6847 | using f_int_acbd uwvz_sub integrable_on_subcbox by blast | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6848 | have f_int_uv: "\<And>x. x \<in> cbox u v \<Longrightarrow> (\<lambda>y. f (x,y)) integrable_on cbox w z" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6849 | using assms continuous_on_subset uwvz_sub | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6850 | by (blast intro: continuous_on_imp_integrable_on_Pair1) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6851 | have 1: "norm (integral (cbox (u,w) (v,z)) f - integral (cbox (u,w) (v,z)) (\<lambda>x. f (t1,t2))) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 6852 | \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX/2" | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 6853 | apply (simp only: integral_diff [symmetric] f_int_uwvz integrable_const) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 6854 | apply (rule order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) | 
| 66505 | 6855 | using cbp \<open>0 < e/content ?CBOX\<close> nf' | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 6856 | apply (auto simp: integrable_diff f_int_uwvz integrable_const) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6857 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6858 | have int_integrable: "(\<lambda>x. integral (cbox w z) (\<lambda>y. f (x, y))) integrable_on cbox u v" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6859 | using assms integral_integrable_2dim continuous_on_subset uwvz_sub(2) by blast | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6860 | have normint_wz: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6861 | "\<And>x. x \<in> cbox u v \<Longrightarrow> | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6862 | norm (integral (cbox w z) (\<lambda>y. f (x, y)) - integral (cbox w z) (\<lambda>y. f (t1, t2))) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 6863 | \<le> e * content (cbox w z) / content (cbox (a, c) (b, d))/2" | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 6864 | apply (simp only: integral_diff [symmetric] f_int_uv integrable_const) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 6865 | apply (rule order_trans [OF integrable_bound [of "e/content ?CBOX/2"]]) | 
| 66505 | 6866 | using cbp \<open>0 < e/content ?CBOX\<close> nf' | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 6867 | apply (auto simp: integrable_diff f_int_uv integrable_const) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6868 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6869 | have "norm (integral (cbox u v) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6870 | (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y)) - integral (cbox w z) (\<lambda>y. f (t1,t2)))) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 6871 | \<le> e * content (cbox w z) / content ?CBOX/2 * content (cbox u v)" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6872 | apply (rule integrable_bound [OF _ _ normint_wz]) | 
| 66505 | 6873 | using cbp \<open>0 < e/content ?CBOX\<close> | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 6874 | apply (auto simp: divide_simps content_pos_le integrable_diff int_integrable integrable_const) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6875 | done | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 6876 | also have "... \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX/2" | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6877 | by (simp add: content_Pair divide_simps) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6878 | finally have 2: "norm (integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y))) - | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6879 | integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (t1,t2)))) | 
| 66532 
8a6ce2d9a9f5
work on indefinite_integral_continuous_left, etc.
 paulson <lp15@cam.ac.uk> parents: 
66524diff
changeset | 6880 | \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX/2" | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61762diff
changeset | 6881 | by (simp only: integral_diff [symmetric] int_integrable integrable_const) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6882 | have norm_xx: "\<lbrakk>x' = y'; norm(x - x') \<le> e/2; norm(y - y') \<le> e/2\<rbrakk> \<Longrightarrow> norm(x - y) \<le> e" for x::'c and y x' y' e | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6883 | using norm_triangle_mono [of "x-y'" "e/2" "y'-y" "e/2"] real_sum_of_halves | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6884 | by (simp add: norm_minus_commute) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6885 | have "norm (integral (cbox (u,w) (v,z)) f - integral (cbox u v) (\<lambda>x. integral (cbox w z) (\<lambda>y. f (x,y)))) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6886 | \<le> e * content (cbox (u,w) (v,z)) / content ?CBOX" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6887 | by (rule norm_xx [OF integral_Pair_const 1 2]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6888 | } note * = this | 
| 66703 | 6889 | have "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) \<le> e" | 
| 6890 | if "\<forall>x\<in>?CBOX. \<forall>x'\<in>?CBOX. norm (x' - x) < k \<longrightarrow> norm (f x' - f x) < e /(2 * content (?CBOX))" "0 < k" for k | |
| 6891 | proof - | |
| 6892 | obtain p where ptag: "p tagged_division_of cbox (a, c) (b, d)" | |
| 6893 | and fine: "(\<lambda>x. ball x k) fine p" | |
| 6894 | using fine_division_exists \<open>0 < k\<close> by blast | |
| 6895 | show ?thesis | |
| 6896 | apply (rule op_acbd [OF division_of_tagged_division [OF ptag]]) | |
| 6897 | using that fine ptag \<open>0 < k\<close> by (auto simp: *) | |
| 6898 | qed | |
| 6899 | then show "norm (integral ?CBOX f - integral (cbox a b) (\<lambda>x. integral (cbox c d) (\<lambda>y. f (x,y)))) \<le> e" | |
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6900 | using compact_uniformly_continuous [OF assms compact_cbox] | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6901 | apply (simp add: uniformly_continuous_on_def dist_norm) | 
| 66406 
f8f4cf0fa42d
cleanup of integral_norm_bound_integral
 paulson <lp15@cam.ac.uk> parents: 
66402diff
changeset | 6902 | apply (drule_tac x="e/2 / content?CBOX" in spec) | 
| 66703 | 6903 | using cbp \<open>0 < e\<close> by (auto simp: zero_less_mult_iff) | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6904 | qed | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6905 | then show ?thesis | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6906 | by simp | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6907 | qed | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6908 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6909 | lemma integral_swap_2dim: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6910 | fixes f :: "['a::euclidean_space, 'b::euclidean_space] \<Rightarrow> 'c::banach" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6911 | assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6912 | shows "integral (cbox (a, c) (b, d)) (\<lambda>(x, y). f x y) = integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6913 | proof - | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6914 | have "((\<lambda>(x, y). f x y) has_integral integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)) (prod.swap ` (cbox (c, a) (d, b)))" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6915 | apply (rule has_integral_twiddle [of 1 prod.swap prod.swap "\<lambda>(x,y). f y x" "integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)", simplified]) | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 6916 | apply (force simp: isCont_swap content_Pair has_integral_integral [symmetric] integrable_continuous swap_continuous assms)+ | 
| 60615 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6917 | done | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6918 | then show ?thesis | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6919 | by force | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6920 | qed | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6921 | |
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6922 | theorem integral_swap_continuous: | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6923 | fixes f :: "['a::euclidean_space, 'b::euclidean_space] \<Rightarrow> 'c::banach" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6924 | assumes "continuous_on (cbox (a,c) (b,d)) (\<lambda>(x,y). f x y)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6925 | shows "integral (cbox a b) (\<lambda>x. integral (cbox c d) (f x)) = | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6926 | integral (cbox c d) (\<lambda>y. integral (cbox a b) (\<lambda>x. f x y))" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6927 | proof - | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6928 | have "integral (cbox a b) (\<lambda>x. integral (cbox c d) (f x)) = integral (cbox (a, c) (b, d)) (\<lambda>(x, y). f x y)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6929 | using integral_prod_continuous [OF assms] by auto | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6930 | also have "... = integral (cbox (c, a) (d, b)) (\<lambda>(x, y). f y x)" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6931 | by (rule integral_swap_2dim [OF assms]) | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6932 | also have "... = integral (cbox c d) (\<lambda>y. integral (cbox a b) (\<lambda>x. f x y))" | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6933 | using integral_prod_continuous [OF swap_continuous] assms | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6934 | by auto | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6935 | finally show ?thesis . | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6936 | qed | 
| 
e5fa1d5d3952
Useful lemmas. The theorem concerning swapping the variables in a double integral.
 paulson <lp15@cam.ac.uk> parents: 
60585diff
changeset | 6937 | |
| 63296 | 6938 | |
| 6939 | subsection \<open>Definite integrals for exponential and power function\<close> | |
| 6940 | ||
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 6941 | lemma has_integral_exp_minus_to_infinity: | 
| 63296 | 6942 | assumes a: "a > 0" | 
| 6943 |   shows   "((\<lambda>x::real. exp (-a*x)) has_integral exp (-a*c)/a) {c..}"
 | |
| 6944 | proof - | |
| 6945 |   define f where "f = (\<lambda>k x. if x \<in> {c..real k} then exp (-a*x) else 0)"
 | |
| 6946 |   {
 | |
| 6947 | fix k :: nat assume k: "of_nat k \<ge> c" | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 6948 | from k a | 
| 63296 | 6949 |       have "((\<lambda>x. exp (-a*x)) has_integral (-exp (-a*real k)/a - (-exp (-a*c)/a))) {c..real k}"
 | 
| 6950 | by (intro fundamental_theorem_of_calculus) | |
| 6951 | (auto intro!: derivative_eq_intros | |
| 6952 | simp: has_field_derivative_iff_has_vector_derivative [symmetric]) | |
| 6953 |     hence "(f k has_integral (exp (-a*c)/a - exp (-a*real k)/a)) {c..}" unfolding f_def
 | |
| 6954 | by (subst has_integral_restrict) simp_all | |
| 6955 | } note has_integral_f = this | |
| 6956 | ||
| 6957 | have [simp]: "f k = (\<lambda>_. 0)" if "of_nat k < c" for k using that by (auto simp: fun_eq_iff f_def) | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 6958 |   have integral_f: "integral {c..} (f k) =
 | 
| 63296 | 6959 | (if real k \<ge> c then exp (-a*c)/a - exp (-a*real k)/a else 0)" | 
| 6960 | for k using integral_unique[OF has_integral_f[of k]] by simp | |
| 6961 | ||
| 6962 |   have A: "(\<lambda>x. exp (-a*x)) integrable_on {c..} \<and>
 | |
| 6963 |              (\<lambda>k. integral {c..} (f k)) \<longlonglongrightarrow> integral {c..} (\<lambda>x. exp (-a*x))"
 | |
| 6964 | proof (intro monotone_convergence_increasing allI ballI) | |
| 6965 | fix k ::nat | |
| 6966 |     have "(\<lambda>x. exp (-a*x)) integrable_on {c..of_real k}" (is ?P)
 | |
| 6967 | unfolding f_def by (auto intro!: continuous_intros integrable_continuous_real) | |
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 6968 |     hence  "(f k) integrable_on {c..of_real k}"
 | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 6969 | by (rule integrable_eq) (simp add: f_def) | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 6970 |     then show "f k integrable_on {c..}"
 | 
| 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 6971 | by (rule integrable_on_superset) (auto simp: f_def) | 
| 63296 | 6972 | next | 
| 6973 |     fix x assume x: "x \<in> {c..}"
 | |
| 6974 |     have "sequentially \<le> principal {nat \<lceil>x\<rceil>..}" unfolding at_top_def by (simp add: Inf_lower)
 | |
| 6975 |     also have "{nat \<lceil>x\<rceil>..} \<subseteq> {k. x \<le> real k}" by auto
 | |
| 6976 |     also have "inf (principal \<dots>) (principal {k. \<not>x \<le> real k}) =
 | |
| 6977 |                  principal ({k. x \<le> real k} \<inter> {k. \<not>x \<le> real k})" by simp
 | |
| 6978 |     also have "{k. x \<le> real k} \<inter> {k. \<not>x \<le> real k} = {}" by blast
 | |
| 6979 |     finally have "inf sequentially (principal {k. \<not>x \<le> real k}) = bot"
 | |
| 6980 | by (simp add: inf.coboundedI1 bot_unique) | |
| 6981 | with x show "(\<lambda>k. f k x) \<longlonglongrightarrow> exp (-a*x)" unfolding f_def | |
| 6982 | by (intro filterlim_If) auto | |
| 6983 | next | |
| 6984 |     have "\<bar>integral {c..} (f k)\<bar> \<le> exp (-a*c)/a" for k
 | |
| 6985 | proof (cases "c > of_nat k") | |
| 6986 | case False | |
| 6987 |       hence "abs (integral {c..} (f k)) = abs (exp (- (a * c)) / a - exp (- (a * real k)) / a)"
 | |
| 6988 | by (simp add: integral_f) | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 6989 | also have "abs (exp (- (a * c)) / a - exp (- (a * real k)) / a) = | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 6990 | exp (- (a * c)) / a - exp (- (a * real k)) / a" | 
| 63296 | 6991 | using False a by (intro abs_of_nonneg) (simp_all add: field_simps) | 
| 6992 | also have "\<dots> \<le> exp (- a * c) / a" using a by simp | |
| 6993 | finally show ?thesis . | |
| 6994 | qed (insert a, simp_all add: integral_f) | |
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 6995 |     thus "bounded (range(\<lambda>k. integral {c..} (f k)))"
 | 
| 65036 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 6996 | by (intro boundedI[of _ "exp (-a*c)/a"]) auto | 
| 63296 | 6997 | qed (auto simp: f_def) | 
| 6998 | ||
| 6999 | from eventually_gt_at_top[of "nat \<lceil>c\<rceil>"] have "eventually (\<lambda>k. of_nat k > c) sequentially" | |
| 7000 | by eventually_elim linarith | |
| 7001 |   hence "eventually (\<lambda>k. exp (-a*c)/a - exp (-a * of_nat k)/a = integral {c..} (f k)) sequentially"
 | |
| 7002 | by eventually_elim (simp add: integral_f) | |
| 7003 | moreover have "(\<lambda>k. exp (-a*c)/a - exp (-a * of_nat k)/a) \<longlonglongrightarrow> exp (-a*c)/a - 0/a" | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7004 | by (intro tendsto_intros filterlim_compose[OF exp_at_bot] | 
| 63296 | 7005 | filterlim_tendsto_neg_mult_at_bot[OF tendsto_const] filterlim_real_sequentially)+ | 
| 7006 | (insert a, simp_all) | |
| 7007 |   ultimately have "(\<lambda>k. integral {c..} (f k)) \<longlonglongrightarrow> exp (-a*c)/a - 0/a"
 | |
| 7008 | by (rule Lim_transform_eventually) | |
| 7009 | from LIMSEQ_unique[OF conjunct2[OF A] this] | |
| 7010 |     have "integral {c..} (\<lambda>x. exp (-a*x)) = exp (-a*c)/a" by simp
 | |
| 7011 | with conjunct1[OF A] show ?thesis | |
| 7012 | by (simp add: has_integral_integral) | |
| 7013 | qed | |
| 7014 | ||
| 7015 | lemma integrable_on_exp_minus_to_infinity: "a > 0 \<Longrightarrow> (\<lambda>x. exp (-a*x) :: real) integrable_on {c..}"
 | |
| 7016 | using has_integral_exp_minus_to_infinity[of a c] unfolding integrable_on_def by blast | |
| 7017 | ||
| 7018 | lemma has_integral_powr_from_0: | |
| 7019 | assumes a: "a > (-1::real)" and c: "c \<ge> 0" | |
| 7020 |   shows   "((\<lambda>x. x powr a) has_integral (c powr (a+1) / (a+1))) {0..c}"
 | |
| 7021 | proof (cases "c = 0") | |
| 7022 | case False | |
| 7023 |   define f where "f = (\<lambda>k x. if x \<in> {inverse (of_nat (Suc k))..c} then x powr a else 0)"
 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7024 | define F where "F = (\<lambda>k. if inverse (of_nat (Suc k)) \<le> c then | 
| 63296 | 7025 | c powr (a+1)/(a+1) - inverse (real (Suc k)) powr (a+1)/(a+1) else 0)" | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7026 | |
| 63296 | 7027 |   {
 | 
| 7028 | fix k :: nat | |
| 7029 |     have "(f k has_integral F k) {0..c}"
 | |
| 7030 | proof (cases "inverse (of_nat (Suc k)) \<le> c") | |
| 7031 | case True | |
| 7032 |       {
 | |
| 7033 | fix x assume x: "x \<ge> inverse (1 + real k)" | |
| 7034 | have "0 < inverse (1 + real k)" by simp | |
| 7035 | also note x | |
| 7036 | finally have "x > 0" . | |
| 7037 | } note x = this | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7038 | hence "((\<lambda>x. x powr a) has_integral c powr (a + 1) / (a + 1) - | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7039 |                inverse (real (Suc k)) powr (a + 1) / (a + 1)) {inverse (real (Suc k))..c}"
 | 
| 63296 | 7040 | using True a by (intro fundamental_theorem_of_calculus) | 
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7041 | (auto intro!: derivative_eq_intros continuous_on_powr' continuous_on_const | 
| 63296 | 7042 | continuous_on_id simp: has_field_derivative_iff_has_vector_derivative [symmetric]) | 
| 7043 | with True show ?thesis unfolding f_def F_def by (subst has_integral_restrict) simp_all | |
| 7044 | next | |
| 7045 | case False | |
| 7046 | thus ?thesis unfolding f_def F_def by (subst has_integral_restrict) auto | |
| 7047 | qed | |
| 7048 | } note has_integral_f = this | |
| 7049 |   have integral_f: "integral {0..c} (f k) = F k" for k
 | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7050 | using has_integral_f[of k] by (rule integral_unique) | 
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7051 | |
| 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7052 |   have A: "(\<lambda>x. x powr a) integrable_on {0..c} \<and>
 | 
| 63296 | 7053 |            (\<lambda>k. integral {0..c} (f k)) \<longlonglongrightarrow> integral {0..c} (\<lambda>x. x powr a)"
 | 
| 7054 | proof (intro monotone_convergence_increasing ballI allI) | |
| 7055 |     fix k from has_integral_f[of k] show "f k integrable_on {0..c}"
 | |
| 7056 | by (auto simp: integrable_on_def) | |
| 7057 | next | |
| 7058 | fix k :: nat and x :: real | |
| 7059 |     {
 | |
| 7060 | assume x: "inverse (real (Suc k)) \<le> x" | |
| 7061 | have "inverse (real (Suc (Suc k))) \<le> inverse (real (Suc k))" by (simp add: field_simps) | |
| 7062 | also note x | |
| 7063 | finally have "inverse (real (Suc (Suc k))) \<le> x" . | |
| 7064 | } | |
| 7065 | thus "f k x \<le> f (Suc k) x" by (auto simp: f_def simp del: of_nat_Suc) | |
| 7066 | next | |
| 7067 |     fix x assume x: "x \<in> {0..c}"
 | |
| 7068 | show "(\<lambda>k. f k x) \<longlonglongrightarrow> x powr a" | |
| 7069 | proof (cases "x = 0") | |
| 7070 | case False | |
| 7071 | with x have "x > 0" by simp | |
| 7072 | from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] | |
| 7073 | have "eventually (\<lambda>k. x powr a = f k x) sequentially" | |
| 7074 | by eventually_elim (insert x, simp add: f_def) | |
| 7075 | moreover have "(\<lambda>_. x powr a) \<longlonglongrightarrow> x powr a" by simp | |
| 7076 | ultimately show ?thesis by (rule Lim_transform_eventually) | |
| 7077 | qed (simp_all add: f_def) | |
| 7078 | next | |
| 7079 |     {
 | |
| 7080 | fix k | |
| 7081 | from a have "F k \<le> c powr (a + 1) / (a + 1)" | |
| 7082 | by (auto simp add: F_def divide_simps) | |
| 63469 
b6900858dcb9
lots of new theorems about differentiable_on, retracts, ANRs, etc.
 paulson <lp15@cam.ac.uk> parents: 
63305diff
changeset | 7083 | also from a have "F k \<ge> 0" | 
| 63296 | 7084 | by (auto simp: F_def divide_simps simp del: of_nat_Suc intro!: powr_mono2) | 
| 7085 | hence "F k = abs (F k)" by simp | |
| 7086 | finally have "abs (F k) \<le> c powr (a + 1) / (a + 1)" . | |
| 7087 | } | |
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 7088 |     thus "bounded (range(\<lambda>k. integral {0..c} (f k)))"
 | 
| 65036 
ab7e11730ad8
Some new lemmas. Existing lemmas modified to use uniform_limit rather than its expansion
 paulson <lp15@cam.ac.uk> parents: 
64911diff
changeset | 7089 | by (intro boundedI[of _ "c powr (a+1) / (a+1)"]) (auto simp: integral_f) | 
| 63296 | 7090 | qed | 
| 7091 | ||
| 7092 | from False c have "c > 0" by simp | |
| 7093 | from order_tendstoD(2)[OF LIMSEQ_inverse_real_of_nat this] | |
| 7094 | have "eventually (\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a+1) / (a+1) = | |
| 7095 |             integral {0..c} (f k)) sequentially"
 | |
| 7096 | by eventually_elim (simp add: integral_f F_def) | |
| 7097 | moreover have "(\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) | |
| 7098 | \<longlonglongrightarrow> c powr (a + 1) / (a + 1) - 0 powr (a + 1) / (a + 1)" | |
| 7099 | using a by (intro tendsto_intros LIMSEQ_inverse_real_of_nat) auto | |
| 7100 | hence "(\<lambda>k. c powr (a + 1) / (a + 1) - inverse (real (Suc k)) powr (a + 1) / (a + 1)) | |
| 7101 | \<longlonglongrightarrow> c powr (a + 1) / (a + 1)" by simp | |
| 7102 |   ultimately have "(\<lambda>k. integral {0..c} (f k)) \<longlonglongrightarrow> c powr (a+1) / (a+1)"
 | |
| 7103 | by (rule Lim_transform_eventually) | |
| 7104 |   with A have "integral {0..c} (\<lambda>x. x powr a) = c powr (a+1) / (a+1)"
 | |
| 7105 | by (blast intro: LIMSEQ_unique) | |
| 7106 | with A show ?thesis by (simp add: has_integral_integral) | |
| 7107 | qed (simp_all add: has_integral_refl) | |
| 7108 | ||
| 7109 | lemma integrable_on_powr_from_0: | |
| 7110 | assumes a: "a > (-1::real)" and c: "c \<ge> 0" | |
| 7111 |   shows   "(\<lambda>x. x powr a) integrable_on {0..c}"
 | |
| 7112 | using has_integral_powr_from_0[OF assms] unfolding integrable_on_def by blast | |
| 7113 | ||
| 63721 | 7114 | lemma has_integral_powr_to_inf: | 
| 7115 | fixes a e :: real | |
| 7116 | assumes "e < -1" "a > 0" | |
| 7117 |   shows   "((\<lambda>x. x powr e) has_integral -(a powr (e + 1)) / (e + 1)) {a..}"
 | |
| 7118 | proof - | |
| 7119 |   define f :: "nat \<Rightarrow> real \<Rightarrow> real" where "f = (\<lambda>n x. if x \<in> {a..n} then x powr e else 0)"
 | |
| 7120 | define F where "F = (\<lambda>x. x powr (e + 1) / (e + 1))" | |
| 7121 | ||
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7122 |   have has_integral_f: "(f n has_integral (F n - F a)) {a..}"
 | 
| 63721 | 7123 | if n: "n \<ge> a" for n :: nat | 
| 7124 | proof - | |
| 7125 |     from n assms have "((\<lambda>x. x powr e) has_integral (F n - F a)) {a..n}"
 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7126 | by (intro fundamental_theorem_of_calculus) (auto intro!: derivative_eq_intros | 
| 63721 | 7127 | simp: has_field_derivative_iff_has_vector_derivative [symmetric] F_def) | 
| 7128 |     hence "(f n has_integral (F n - F a)) {a..n}"
 | |
| 7129 | by (rule has_integral_eq [rotated]) (simp add: f_def) | |
| 7130 |     thus "(f n has_integral (F n - F a)) {a..}"
 | |
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 7131 | by (rule has_integral_on_superset) (auto simp: f_def) | 
| 63721 | 7132 | qed | 
| 7133 |   have integral_f: "integral {a..} (f n) = (if n \<ge> a then F n - F a else 0)" for n :: nat
 | |
| 7134 | proof (cases "n \<ge> a") | |
| 7135 | case True | |
| 7136 | with has_integral_f[OF this] show ?thesis by (simp add: integral_unique) | |
| 7137 | next | |
| 7138 | case False | |
| 7139 |     have "(f n has_integral 0) {a}" by (rule has_integral_refl)
 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7140 |     hence "(f n has_integral 0) {a..}"
 | 
| 66164 
2d79288b042c
New theorems and much tidying up of the old ones
 paulson <lp15@cam.ac.uk> parents: 
66154diff
changeset | 7141 | by (rule has_integral_on_superset) (insert False, simp_all add: f_def) | 
| 63721 | 7142 | with False show ?thesis by (simp add: integral_unique) | 
| 7143 | qed | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7144 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7145 |   have *: "(\<lambda>x. x powr e) integrable_on {a..} \<and>
 | 
| 63721 | 7146 |            (\<lambda>n. integral {a..} (f n)) \<longlonglongrightarrow> integral {a..} (\<lambda>x. x powr e)"
 | 
| 7147 | proof (intro monotone_convergence_increasing allI ballI) | |
| 7148 | fix n :: nat | |
| 7149 |     from assms have "(\<lambda>x. x powr e) integrable_on {a..n}"
 | |
| 7150 | by (auto intro!: integrable_continuous_real continuous_intros) | |
| 7151 |     hence "f n integrable_on {a..n}"
 | |
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 7152 | by (rule integrable_eq) (auto simp: f_def) | 
| 63721 | 7153 |     thus "f n integrable_on {a..}"
 | 
| 66552 
507a42c0a0ff
last-minute integration unscrambling
 paulson <lp15@cam.ac.uk> parents: 
66539diff
changeset | 7154 | by (rule integrable_on_superset) (auto simp: f_def) | 
| 63721 | 7155 | next | 
| 7156 | fix n :: nat and x :: real | |
| 7157 | show "f n x \<le> f (Suc n) x" by (auto simp: f_def) | |
| 7158 | next | |
| 7159 |     fix x :: real assume x: "x \<in> {a..}"
 | |
| 7160 | from filterlim_real_sequentially | |
| 7161 | have "eventually (\<lambda>n. real n \<ge> x) at_top" | |
| 7162 | by (simp add: filterlim_at_top) | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7163 | with x have "eventually (\<lambda>n. f n x = x powr e) at_top" | 
| 63721 | 7164 | by (auto elim!: eventually_mono simp: f_def) | 
| 7165 | thus "(\<lambda>n. f n x) \<longlonglongrightarrow> x powr e" by (simp add: Lim_eventually) | |
| 7166 | next | |
| 7167 |     have "norm (integral {a..} (f n)) \<le> -F a" for n :: nat
 | |
| 7168 | proof (cases "n \<ge> a") | |
| 7169 | case True | |
| 7170 | with assms have "a powr (e + 1) \<ge> n powr (e + 1)" | |
| 7171 | by (intro powr_mono2') simp_all | |
| 7172 | with assms show ?thesis by (auto simp: divide_simps F_def integral_f) | |
| 7173 | qed (insert assms, simp add: integral_f F_def divide_simps) | |
| 66408 
46cfd348c373
general rationalisation of Analysis
 paulson <lp15@cam.ac.uk> parents: 
66406diff
changeset | 7174 |     thus "bounded (range(\<lambda>k. integral {a..} (f k)))"
 | 
| 63721 | 7175 | unfolding bounded_iff by (intro exI[of _ "-F a"]) auto | 
| 7176 | qed | |
| 7177 | ||
| 7178 | from filterlim_real_sequentially | |
| 7179 | have "eventually (\<lambda>n. real n \<ge> a) at_top" | |
| 7180 | by (simp add: filterlim_at_top) | |
| 7181 |   hence "eventually (\<lambda>n. F n - F a = integral {a..} (f n)) at_top"
 | |
| 7182 | by eventually_elim (simp add: integral_f) | |
| 7183 | moreover have "(\<lambda>n. F n - F a) \<longlonglongrightarrow> 0 / (e + 1) - F a" unfolding F_def | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7184 | by (insert assms, (rule tendsto_intros filterlim_compose[OF tendsto_neg_powr] | 
| 63721 | 7185 | filterlim_ident filterlim_real_sequentially | simp)+) | 
| 7186 | hence "(\<lambda>n. F n - F a) \<longlonglongrightarrow> -F a" by simp | |
| 7187 |   ultimately have "(\<lambda>n. integral {a..} (f n)) \<longlonglongrightarrow> -F a" by (rule Lim_transform_eventually)
 | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7188 | from conjunct2[OF *] and this | 
| 63721 | 7189 |     have "integral {a..} (\<lambda>x. x powr e) = -F a" by (rule LIMSEQ_unique)
 | 
| 7190 | with conjunct1[OF *] show ?thesis | |
| 7191 | by (simp add: has_integral_integral F_def) | |
| 7192 | qed | |
| 7193 | ||
| 7194 | lemma has_integral_inverse_power_to_inf: | |
| 7195 | fixes a :: real and n :: nat | |
| 7196 | assumes "n > 1" "a > 0" | |
| 7197 |   shows   "((\<lambda>x. 1 / x ^ n) has_integral 1 / (real (n - 1) * a ^ (n - 1))) {a..}"
 | |
| 7198 | proof - | |
| 7199 | from assms have "real_of_int (-int n) < -1" by simp | |
| 7200 | note has_integral_powr_to_inf[OF this \<open>a > 0\<close>] | |
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7201 | also have "- (a powr (real_of_int (- int n) + 1)) / (real_of_int (- int n) + 1) = | 
| 63721 | 7202 | 1 / (real (n - 1) * a powr (real (n - 1)))" using assms | 
| 7203 | by (simp add: divide_simps powr_add [symmetric] of_nat_diff) | |
| 7204 | also from assms have "a powr (real (n - 1)) = a ^ (n - 1)" | |
| 7205 | by (intro powr_realpow) | |
| 7206 | finally show ?thesis | |
| 7207 | by (rule has_integral_eq [rotated]) | |
| 7208 | (insert assms, simp_all add: powr_minus powr_realpow divide_simps) | |
| 7209 | qed | |
| 7210 | ||
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7211 | subsubsection \<open>Adaption to ordered Euclidean spaces and the Cartesian Euclidean space\<close> | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7212 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7213 | lemma integral_component_eq_cart[simp]: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7214 | fixes f :: "'n::euclidean_space \<Rightarrow> real^'m" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7215 | assumes "f integrable_on s" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7216 | shows "integral s (\<lambda>x. f x $ k) = integral s f $ k" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7217 | using integral_linear[OF assms(1) bounded_linear_component_cart,unfolded o_def] . | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7218 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7219 | lemma content_closed_interval: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7220 | fixes a :: "'a::ordered_euclidean_space" | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7221 | assumes "a \<le> b" | 
| 66402 | 7222 |   shows "content {a..b} = (\<Prod>i\<in>Basis. b\<bullet>i - a\<bullet>i)"
 | 
| 66409 | 7223 | using content_cbox[of a b] assms by (simp add: cbox_interval eucl_le[where 'a='a]) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7224 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7225 | lemma integrable_const_ivl[intro]: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7226 | fixes a::"'a::ordered_euclidean_space" | 
| 66402 | 7227 |   shows "(\<lambda>x. c) integrable_on {a..b}"
 | 
| 66409 | 7228 | unfolding cbox_interval[symmetric] by (rule integrable_const) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7229 | |
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7230 | lemma integrable_on_subinterval: | 
| 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7231 | fixes f :: "'n::ordered_euclidean_space \<Rightarrow> 'a::banach" | 
| 66409 | 7232 |   assumes "f integrable_on s" "{a..b} \<subseteq> s"
 | 
| 66402 | 7233 |   shows "f integrable_on {a..b}"
 | 
| 66409 | 7234 | using integrable_on_subcbox[of f s a b] assms by (simp add: cbox_interval) | 
| 63886 
685fb01256af
move Henstock-Kurzweil integration after Lebesgue_Measure; replace content by abbreviation measure lborel
 hoelzl parents: 
63721diff
changeset | 7235 | |
| 35173 
9b24bfca8044
Renamed Multivariate-Analysis/Integration to Multivariate-Analysis/Integration_MV to avoid name clash with Integration.
 hoelzl parents: 
35172diff
changeset | 7236 | end |