| author | haftmann | 
| Mon, 26 Oct 2020 11:28:43 +0000 | |
| changeset 72508 | c89d8e8bd8c7 | 
| parent 68634 | db0980691ef4 | 
| permissions | -rw-r--r-- | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31366diff
changeset | 1 | (* Author: Jacques D. Fleuriot, University of Edinburgh | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 2 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | 
| 35328 | 3 | |
| 63627 | 4 | Replaced by ~~/src/HOL/Analysis/Henstock_Kurzweil_Integration and | 
| 5 | Bochner_Integration. | |
| 13958 | 6 | *) | 
| 7 | ||
| 61343 | 8 | section\<open>Theory of Integration on real intervals\<close> | 
| 35328 | 9 | |
| 10 | theory Gauge_Integration | |
| 11 | imports Complex_Main | |
| 12 | begin | |
| 13 | ||
| 61343 | 14 | text \<open> | 
| 13958 | 15 | |
| 35328 | 16 | \textbf{Attention}: This theory defines the Integration on real
 | 
| 17 | intervals. This is just a example theory for historical / expository interests. | |
| 18 | A better replacement is found in the Multivariate Analysis library. This defines | |
| 19 | the gauge integral on real vector spaces and in the Real Integral theory | |
| 20 | is a specialization to the integral on arbitrary real intervals. The | |
| 21 | Multivariate Analysis package also provides a better support for analysis on | |
| 22 | integrals. | |
| 23 | ||
| 61343 | 24 | \<close> | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 25 | |
| 61343 | 26 | text\<open>We follow John Harrison in formalizing the Gauge integral.\<close> | 
| 13958 | 27 | |
| 61343 | 28 | subsection \<open>Gauges\<close> | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 29 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20792diff
changeset | 30 | definition | 
| 31253 | 31 | gauge :: "[real set, real => real] => bool" where | 
| 37765 | 32 | "gauge E g = (\<forall>x\<in>E. 0 < g(x))" | 
| 13958 | 33 | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 34 | |
| 61343 | 35 | subsection \<open>Gauge-fine divisions\<close> | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 36 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 37 | inductive | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 38 | fine :: "[real \<Rightarrow> real, real \<times> real, (real \<times> real \<times> real) list] \<Rightarrow> bool" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 39 | for | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 40 | \<delta> :: "real \<Rightarrow> real" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 41 | where | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 42 | fine_Nil: | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 43 | "fine \<delta> (a, a) []" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 44 | | fine_Cons: | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 45 | "\<lbrakk>fine \<delta> (b, c) D; a < b; a \<le> x; x \<le> b; b - a < \<delta> x\<rbrakk> | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 46 | \<Longrightarrow> fine \<delta> (a, c) ((a, x, b) # D)" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 47 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 48 | lemmas fine_induct [induct set: fine] = | 
| 61424 
c3658c18b7bc
prod_case as canonical name for product type eliminator
 haftmann parents: 
61343diff
changeset | 49 | fine.induct [of "\<delta>" "(a,b)" "D" "case_prod P", unfolded split_conv] for \<delta> a b D P | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 50 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 51 | lemma fine_single: | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 52 | "\<lbrakk>a < b; a \<le> x; x \<le> b; b - a < \<delta> x\<rbrakk> \<Longrightarrow> fine \<delta> (a, b) [(a, x, b)]" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 53 | by (rule fine_Cons [OF fine_Nil]) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 54 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 55 | lemma fine_append: | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 56 | "\<lbrakk>fine \<delta> (a, b) D; fine \<delta> (b, c) D'\<rbrakk> \<Longrightarrow> fine \<delta> (a, c) (D @ D')" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 57 | by (induct set: fine, simp, simp add: fine_Cons) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 58 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 59 | lemma fine_imp_le: "fine \<delta> (a, b) D \<Longrightarrow> a \<le> b" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 60 | by (induct set: fine, simp_all) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 61 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 62 | lemma nonempty_fine_imp_less: "\<lbrakk>fine \<delta> (a, b) D; D \<noteq> []\<rbrakk> \<Longrightarrow> a < b" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 63 | apply (induct set: fine, simp) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 64 | apply (drule fine_imp_le, simp) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 65 | done | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 66 | |
| 35441 | 67 | lemma fine_Nil_iff: "fine \<delta> (a, b) [] \<longleftrightarrow> a = b" | 
| 68 | by (auto elim: fine.cases intro: fine.intros) | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 69 | |
| 35441 | 70 | lemma fine_same_iff: "fine \<delta> (a, a) D \<longleftrightarrow> D = []" | 
| 71 | proof | |
| 72 | assume "fine \<delta> (a, a) D" thus "D = []" | |
| 73 | by (metis nonempty_fine_imp_less less_irrefl) | |
| 74 | next | |
| 75 | assume "D = []" thus "fine \<delta> (a, a) D" | |
| 76 | by (simp add: fine_Nil) | |
| 77 | qed | |
| 78 | ||
| 79 | lemma empty_fine_imp_eq: "\<lbrakk>fine \<delta> (a, b) D; D = []\<rbrakk> \<Longrightarrow> a = b" | |
| 80 | by (simp add: fine_Nil_iff) | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 81 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 82 | lemma mem_fine: | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 83 | "\<lbrakk>fine \<delta> (a, b) D; (u, x, v) \<in> set D\<rbrakk> \<Longrightarrow> u < v \<and> u \<le> x \<and> x \<le> v" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 84 | by (induct set: fine, simp, force) | 
| 13958 | 85 | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 86 | lemma mem_fine2: "\<lbrakk>fine \<delta> (a, b) D; (u, z, v) \<in> set D\<rbrakk> \<Longrightarrow> a \<le> u \<and> v \<le> b" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 87 | apply (induct arbitrary: z u v set: fine, auto) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 88 | apply (simp add: fine_imp_le) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 89 | apply (erule order_trans [OF less_imp_le], simp) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 90 | done | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 91 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 92 | lemma mem_fine3: "\<lbrakk>fine \<delta> (a, b) D; (u, z, v) \<in> set D\<rbrakk> \<Longrightarrow> v - u < \<delta> z" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 93 | by (induct arbitrary: z u v set: fine) auto | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 94 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 95 | lemma BOLZANO: | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 96 | fixes P :: "real \<Rightarrow> real \<Rightarrow> bool" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 97 | assumes 1: "a \<le> b" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 98 | assumes 2: "\<And>a b c. \<lbrakk>P a b; P b c; a \<le> b; b \<le> c\<rbrakk> \<Longrightarrow> P a c" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 99 | assumes 3: "\<And>x. \<exists>d>0. \<forall>a b. a \<le> x & x \<le> b & (b-a) < d \<longrightarrow> P a b" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 100 | shows "P a b" | 
| 51480 
3793c3a11378
move connected to HOL image; used to show intermediate value theorem
 hoelzl parents: 
49962diff
changeset | 101 | using 1 2 3 by (rule Bolzano) | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 102 | |
| 61343 | 103 | text\<open>We can always find a division that is fine wrt any gauge\<close> | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 104 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 105 | lemma fine_exists: | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 106 |   assumes "a \<le> b" and "gauge {a..b} \<delta>" shows "\<exists>D. fine \<delta> (a, b) D"
 | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 107 | proof - | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 108 |   {
 | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 109 | fix u v :: real assume "u \<le> v" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 110 | have "a \<le> u \<Longrightarrow> v \<le> b \<Longrightarrow> \<exists>D. fine \<delta> (u, v) D" | 
| 61343 | 111 | apply (induct u v rule: BOLZANO, rule \<open>u \<le> v\<close>) | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 112 | apply (simp, fast intro: fine_append) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 113 | apply (case_tac "a \<le> x \<and> x \<le> b") | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 114 | apply (rule_tac x="\<delta> x" in exI) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 115 | apply (rule conjI) | 
| 61343 | 116 |         apply (simp add: \<open>gauge {a..b} \<delta>\<close> [unfolded gauge_def])
 | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 117 | apply (clarify, rename_tac u v) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 118 | apply (case_tac "u = v") | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 119 | apply (fast intro: fine_Nil) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 120 | apply (subgoal_tac "u < v", fast intro: fine_single, simp) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 121 | apply (rule_tac x="1" in exI, clarsimp) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 122 | done | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 123 | } | 
| 61343 | 124 | with \<open>a \<le> b\<close> show ?thesis by auto | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 125 | qed | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 126 | |
| 31364 | 127 | lemma fine_covers_all: | 
| 128 | assumes "fine \<delta> (a, c) D" and "a < x" and "x \<le> c" | |
| 129 | shows "\<exists> N < length D. \<forall> d t e. D ! N = (d,t,e) \<longrightarrow> d < x \<and> x \<le> e" | |
| 130 | using assms | |
| 131 | proof (induct set: fine) | |
| 132 | case (2 b c D a t) | |
| 133 | thus ?case | |
| 134 | proof (cases "b < x") | |
| 135 | case True | |
| 136 | with 2 obtain N where *: "N < length D" | |
| 63060 | 137 | and **: "D ! N = (d,t,e) \<Longrightarrow> d < x \<and> x \<le> e" for d t e by auto | 
| 31364 | 138 | hence "Suc N < length ((a,t,b)#D) \<and> | 
| 139 | (\<forall> d t' e. ((a,t,b)#D) ! Suc N = (d,t',e) \<longrightarrow> d < x \<and> x \<le> e)" by auto | |
| 140 | thus ?thesis by auto | |
| 141 | next | |
| 142 | case False with 2 | |
| 143 | have "0 < length ((a,t,b)#D) \<and> | |
| 144 | (\<forall> d t' e. ((a,t,b)#D) ! 0 = (d,t',e) \<longrightarrow> d < x \<and> x \<le> e)" by auto | |
| 145 | thus ?thesis by auto | |
| 146 | qed | |
| 147 | qed auto | |
| 148 | ||
| 149 | lemma fine_append_split: | |
| 150 | assumes "fine \<delta> (a,b) D" and "D2 \<noteq> []" and "D = D1 @ D2" | |
| 151 | shows "fine \<delta> (a,fst (hd D2)) D1" (is "?fine1") | |
| 152 | and "fine \<delta> (fst (hd D2), b) D2" (is "?fine2") | |
| 153 | proof - | |
| 154 | from assms | |
| 155 | have "?fine1 \<and> ?fine2" | |
| 156 | proof (induct arbitrary: D1 D2) | |
| 157 | case (2 b c D a' x D1 D2) | |
| 158 | note induct = this | |
| 159 | ||
| 160 | thus ?case | |
| 161 | proof (cases D1) | |
| 162 | case Nil | |
| 163 | hence "fst (hd D2) = a'" using 2 by auto | |
| 61343 | 164 | with fine_Cons[OF \<open>fine \<delta> (b,c) D\<close> induct(3,4,5)] Nil induct | 
| 31364 | 165 | show ?thesis by (auto intro: fine_Nil) | 
| 166 | next | |
| 167 | case (Cons d1 D1') | |
| 61343 | 168 | with induct(2)[OF \<open>D2 \<noteq> []\<close>, of D1'] induct(8) | 
| 31364 | 169 | have "fine \<delta> (b, fst (hd D2)) D1'" and "fine \<delta> (fst (hd D2), c) D2" and | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31366diff
changeset | 170 | "d1 = (a', x, b)" by auto | 
| 31364 | 171 | with fine_Cons[OF this(1) induct(3,4,5), OF induct(6)] Cons | 
| 172 | show ?thesis by auto | |
| 173 | qed | |
| 174 | qed auto | |
| 175 | thus ?fine1 and ?fine2 by auto | |
| 176 | qed | |
| 177 | ||
| 178 | lemma fine_\<delta>_expand: | |
| 179 | assumes "fine \<delta> (a,b) D" | |
| 35441 | 180 | and "\<And>x. a \<le> x \<Longrightarrow> x \<le> b \<Longrightarrow> \<delta> x \<le> \<delta>' x" | 
| 31364 | 181 | shows "fine \<delta>' (a,b) D" | 
| 182 | using assms proof induct | |
| 183 | case 1 show ?case by (rule fine_Nil) | |
| 184 | next | |
| 185 | case (2 b c D a x) | |
| 186 | show ?case | |
| 187 | proof (rule fine_Cons) | |
| 188 | show "fine \<delta>' (b,c) D" using 2 by auto | |
| 61343 | 189 | from fine_imp_le[OF 2(1)] 2(6) \<open>x \<le> b\<close> | 
| 31364 | 190 | show "b - a < \<delta>' x" | 
| 61343 | 191 | using 2(7)[OF \<open>a \<le> x\<close>] by auto | 
| 31364 | 192 | qed (auto simp add: 2) | 
| 193 | qed | |
| 194 | ||
| 195 | lemma fine_single_boundaries: | |
| 196 | assumes "fine \<delta> (a,b) D" and "D = [(d, t, e)]" | |
| 197 | shows "a = d \<and> b = e" | |
| 198 | using assms proof induct | |
| 199 | case (2 b c D a x) | |
| 200 | hence "D = []" and "a = d" and "b = e" by auto | |
| 201 | moreover | |
| 61343 | 202 | from \<open>fine \<delta> (b,c) D\<close> \<open>D = []\<close> have "b = c" | 
| 31364 | 203 | by (rule empty_fine_imp_eq) | 
| 204 | ultimately show ?case by simp | |
| 205 | qed auto | |
| 206 | ||
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63627diff
changeset | 207 | lemma fine_sum_list_eq_diff: | 
| 35328 | 208 | fixes f :: "real \<Rightarrow> real" | 
| 209 | shows "fine \<delta> (a, b) D \<Longrightarrow> (\<Sum>(u, x, v)\<leftarrow>D. f v - f u) = f b - f a" | |
| 210 | by (induct set: fine) simp_all | |
| 211 | ||
| 61343 | 212 | text\<open>Lemmas about combining gauges\<close> | 
| 35328 | 213 | |
| 214 | lemma gauge_min: | |
| 215 | "[| gauge(E) g1; gauge(E) g2 |] | |
| 216 | ==> gauge(E) (%x. min (g1(x)) (g2(x)))" | |
| 217 | by (simp add: gauge_def) | |
| 218 | ||
| 219 | lemma fine_min: | |
| 220 | "fine (%x. min (g1(x)) (g2(x))) (a,b) D | |
| 221 | ==> fine(g1) (a,b) D & fine(g2) (a,b) D" | |
| 222 | apply (erule fine.induct) | |
| 223 | apply (simp add: fine_Nil) | |
| 224 | apply (simp add: fine_Cons) | |
| 225 | done | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 226 | |
| 61343 | 227 | subsection \<open>Riemann sum\<close> | 
| 13958 | 228 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20792diff
changeset | 229 | definition | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 230 | rsum :: "[(real \<times> real \<times> real) list, real \<Rightarrow> real] \<Rightarrow> real" where | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 231 | "rsum D f = (\<Sum>(u, x, v)\<leftarrow>D. f x * (v - u))" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 232 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 233 | lemma rsum_Nil [simp]: "rsum [] f = 0" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 234 | unfolding rsum_def by simp | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 235 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 236 | lemma rsum_Cons [simp]: "rsum ((u, x, v) # D) f = f x * (v - u) + rsum D f" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 237 | unfolding rsum_def by simp | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 238 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 239 | lemma rsum_zero [simp]: "rsum D (\<lambda>x. 0) = 0" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 240 | by (induct D, auto) | 
| 13958 | 241 | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 242 | lemma rsum_left_distrib: "rsum D f * c = rsum D (\<lambda>x. f x * c)" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 243 | by (induct D, auto simp add: algebra_simps) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 244 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 245 | lemma rsum_right_distrib: "c * rsum D f = rsum D (\<lambda>x. c * f x)" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 246 | by (induct D, auto simp add: algebra_simps) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 247 | |
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 248 | lemma rsum_add: "rsum D (\<lambda>x. f x + g x) = rsum D f + rsum D g" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 249 | by (induct D, auto simp add: algebra_simps) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 250 | |
| 31364 | 251 | lemma rsum_append: "rsum (D1 @ D2) f = rsum D1 f + rsum D2 f" | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63627diff
changeset | 252 | unfolding rsum_def map_append sum_list_append .. | 
| 31364 | 253 | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 254 | |
| 61343 | 255 | subsection \<open>Gauge integrability (definite)\<close> | 
| 13958 | 256 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20792diff
changeset | 257 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
20792diff
changeset | 258 | Integral :: "[(real*real),real=>real,real] => bool" where | 
| 37765 | 259 | "Integral = (%(a,b) f k. \<forall>e > 0. | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 260 |                                (\<exists>\<delta>. gauge {a .. b} \<delta> &
 | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 261 | (\<forall>D. fine \<delta> (a,b) D --> | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 262 | \<bar>rsum D f - k\<bar> < e)))" | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 263 | |
| 35441 | 264 | lemma Integral_eq: | 
| 265 | "Integral (a, b) f k \<longleftrightarrow> | |
| 266 |     (\<forall>e>0. \<exists>\<delta>. gauge {a..b} \<delta> \<and> (\<forall>D. fine \<delta> (a,b) D \<longrightarrow> \<bar>rsum D f - k\<bar> < e))"
 | |
| 267 | unfolding Integral_def by simp | |
| 268 | ||
| 269 | lemma IntegralI: | |
| 270 | assumes "\<And>e. 0 < e \<Longrightarrow> | |
| 271 |     \<exists>\<delta>. gauge {a..b} \<delta> \<and> (\<forall>D. fine \<delta> (a, b) D \<longrightarrow> \<bar>rsum D f - k\<bar> < e)"
 | |
| 272 | shows "Integral (a, b) f k" | |
| 273 | using assms unfolding Integral_def by auto | |
| 274 | ||
| 275 | lemma IntegralE: | |
| 276 | assumes "Integral (a, b) f k" and "0 < e" | |
| 277 |   obtains \<delta> where "gauge {a..b} \<delta>" and "\<forall>D. fine \<delta> (a, b) D \<longrightarrow> \<bar>rsum D f - k\<bar> < e"
 | |
| 278 | using assms unfolding Integral_def by auto | |
| 279 | ||
| 31252 | 280 | lemma Integral_def2: | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 281 |   "Integral = (%(a,b) f k. \<forall>e>0. (\<exists>\<delta>. gauge {a..b} \<delta> &
 | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 282 | (\<forall>D. fine \<delta> (a,b) D --> | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 283 | \<bar>rsum D f - k\<bar> \<le> e)))" | 
| 31252 | 284 | unfolding Integral_def | 
| 285 | apply (safe intro!: ext) | |
| 286 | apply (fast intro: less_imp_le) | |
| 287 | apply (drule_tac x="e/2" in spec) | |
| 288 | apply force | |
| 289 | done | |
| 290 | ||
| 61343 | 291 | text\<open>The integral is unique if it exists\<close> | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 292 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 293 | lemma Integral_unique: | 
| 35441 | 294 | assumes le: "a \<le> b" | 
| 295 | assumes 1: "Integral (a, b) f k1" | |
| 296 | assumes 2: "Integral (a, b) f k2" | |
| 297 | shows "k1 = k2" | |
| 298 | proof (rule ccontr) | |
| 299 | assume "k1 \<noteq> k2" | |
| 300 | hence e: "0 < \<bar>k1 - k2\<bar> / 2" by simp | |
| 301 |   obtain d1 where "gauge {a..b} d1" and
 | |
| 302 | d1: "\<forall>D. fine d1 (a, b) D \<longrightarrow> \<bar>rsum D f - k1\<bar> < \<bar>k1 - k2\<bar> / 2" | |
| 303 | using 1 e by (rule IntegralE) | |
| 304 |   obtain d2 where "gauge {a..b} d2" and
 | |
| 305 | d2: "\<forall>D. fine d2 (a, b) D \<longrightarrow> \<bar>rsum D f - k2\<bar> < \<bar>k1 - k2\<bar> / 2" | |
| 306 | using 2 e by (rule IntegralE) | |
| 307 |   have "gauge {a..b} (\<lambda>x. min (d1 x) (d2 x))"
 | |
| 61343 | 308 |     using \<open>gauge {a..b} d1\<close> and \<open>gauge {a..b} d2\<close>
 | 
| 35441 | 309 | by (rule gauge_min) | 
| 310 | then obtain D where "fine (\<lambda>x. min (d1 x) (d2 x)) (a, b) D" | |
| 311 | using fine_exists [OF le] by fast | |
| 312 | hence "fine d1 (a, b) D" and "fine d2 (a, b) D" | |
| 313 | by (auto dest: fine_min) | |
| 314 | hence "\<bar>rsum D f - k1\<bar> < \<bar>k1 - k2\<bar> / 2" and "\<bar>rsum D f - k2\<bar> < \<bar>k1 - k2\<bar> / 2" | |
| 315 | using d1 d2 by simp_all | |
| 316 | hence "\<bar>rsum D f - k1\<bar> + \<bar>rsum D f - k2\<bar> < \<bar>k1 - k2\<bar> / 2 + \<bar>k1 - k2\<bar> / 2" | |
| 317 | by (rule add_strict_mono) | |
| 318 | thus False by auto | |
| 319 | qed | |
| 320 | ||
| 321 | lemma Integral_zero: "Integral(a,a) f 0" | |
| 322 | apply (rule IntegralI) | |
| 323 | apply (rule_tac x = "\<lambda>x. 1" in exI) | |
| 324 | apply (simp add: fine_same_iff gauge_def) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 325 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 326 | |
| 35441 | 327 | lemma Integral_same_iff [simp]: "Integral (a, a) f k \<longleftrightarrow> k = 0" | 
| 328 | by (auto intro: Integral_zero Integral_unique) | |
| 329 | ||
| 330 | lemma Integral_zero_fun: "Integral (a,b) (\<lambda>x. 0) 0" | |
| 331 | apply (rule IntegralI) | |
| 332 | apply (rule_tac x="\<lambda>x. 1" in exI, simp add: gauge_def) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 333 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 334 | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 335 | lemma fine_rsum_const: "fine \<delta> (a,b) D \<Longrightarrow> rsum D (\<lambda>x. c) = (c * (b - a))" | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 336 | unfolding rsum_def | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 337 | by (induct set: fine, auto simp add: algebra_simps) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 338 | |
| 35441 | 339 | lemma Integral_mult_const: "a \<le> b \<Longrightarrow> Integral(a,b) (\<lambda>x. c) (c * (b - a))" | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 340 | apply (cases "a = b", simp) | 
| 35441 | 341 | apply (rule IntegralI) | 
| 342 | apply (rule_tac x = "\<lambda>x. b - a" in exI) | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 343 | apply (rule conjI, simp add: gauge_def) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 344 | apply (clarify) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 345 | apply (subst fine_rsum_const, assumption, simp) | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 346 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 347 | |
| 35441 | 348 | lemma Integral_eq_diff_bounds: "a \<le> b \<Longrightarrow> Integral(a,b) (\<lambda>x. 1) (b - a)" | 
| 349 | using Integral_mult_const [of a b 1] by simp | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 350 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 351 | lemma Integral_mult: | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 352 | "[| a \<le> b; Integral(a,b) f k |] ==> Integral(a,b) (%x. c * f x) (c * k)" | 
| 35441 | 353 | apply (auto simp add: order_le_less) | 
| 354 | apply (cases "c = 0", simp add: Integral_zero_fun) | |
| 355 | apply (rule IntegralI) | |
| 56541 | 356 | apply (erule_tac e="e / \<bar>c\<bar>" in IntegralE, simp) | 
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 357 | apply (rule_tac x="\<delta>" in exI, clarify) | 
| 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 358 | apply (drule_tac x="D" in spec, clarify) | 
| 31257 | 359 | apply (simp add: pos_less_divide_eq abs_mult [symmetric] | 
| 360 | algebra_simps rsum_right_distrib) | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 361 | done | 
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 362 | |
| 31364 | 363 | lemma Integral_add: | 
| 364 | assumes "Integral (a, b) f x1" | |
| 365 | assumes "Integral (b, c) f x2" | |
| 366 | assumes "a \<le> b" and "b \<le> c" | |
| 367 | shows "Integral (a, c) f (x1 + x2)" | |
| 35441 | 368 | proof (cases "a < b \<and> b < c", rule IntegralI) | 
| 31364 | 369 | fix \<epsilon> :: real assume "0 < \<epsilon>" | 
| 370 | hence "0 < \<epsilon> / 2" by auto | |
| 371 | ||
| 372 | assume "a < b \<and> b < c" | |
| 373 | hence "a < b" and "b < c" by auto | |
| 374 | ||
| 375 |   obtain \<delta>1 where \<delta>1_gauge: "gauge {a..b} \<delta>1"
 | |
| 63060 | 376 | and I1: "fine \<delta>1 (a,b) D \<Longrightarrow> \<bar> rsum D f - x1 \<bar> < (\<epsilon> / 2)" for D | 
| 61343 | 377 | using IntegralE [OF \<open>Integral (a, b) f x1\<close> \<open>0 < \<epsilon>/2\<close>] by auto | 
| 31364 | 378 | |
| 379 |   obtain \<delta>2 where \<delta>2_gauge: "gauge {b..c} \<delta>2"
 | |
| 63060 | 380 | and I2: "fine \<delta>2 (b,c) D \<Longrightarrow> \<bar> rsum D f - x2 \<bar> < (\<epsilon> / 2)" for D | 
| 61343 | 381 | using IntegralE [OF \<open>Integral (b, c) f x2\<close> \<open>0 < \<epsilon>/2\<close>] by auto | 
| 31364 | 382 | |
| 63040 | 383 | define \<delta> where "\<delta> x = | 
| 384 | (if x < b then min (\<delta>1 x) (b - x) | |
| 385 | else if x = b then min (\<delta>1 b) (\<delta>2 b) | |
| 386 | else min (\<delta>2 x) (x - b))" for x | |
| 31364 | 387 | |
| 388 |   have "gauge {a..c} \<delta>"
 | |
| 389 | using \<delta>1_gauge \<delta>2_gauge unfolding \<delta>_def gauge_def by auto | |
| 35441 | 390 | |
| 31364 | 391 |   moreover {
 | 
| 392 | fix D :: "(real \<times> real \<times> real) list" | |
| 393 | assume fine: "fine \<delta> (a,c) D" | |
| 61343 | 394 | from fine_covers_all[OF this \<open>a < b\<close> \<open>b \<le> c\<close>] | 
| 31364 | 395 | obtain N where "N < length D" | 
| 396 | and *: "\<forall> d t e. D ! N = (d, t, e) \<longrightarrow> d < b \<and> b \<le> e" | |
| 397 | by auto | |
| 398 | obtain d t e where D_eq: "D ! N = (d, t, e)" by (cases "D!N", auto) | |
| 399 | with * have "d < b" and "b \<le> e" by auto | |
| 400 | have in_D: "(d, t, e) \<in> set D" | |
| 61343 | 401 | using D_eq[symmetric] using \<open>N < length D\<close> by auto | 
| 31364 | 402 | |
| 403 | from mem_fine[OF fine in_D] | |
| 404 | have "d < e" and "d \<le> t" and "t \<le> e" by auto | |
| 405 | ||
| 406 | have "t = b" | |
| 407 | proof (rule ccontr) | |
| 408 | assume "t \<noteq> b" | |
| 61343 | 409 | with mem_fine3[OF fine in_D] \<open>b \<le> e\<close> \<open>d \<le> t\<close> \<open>t \<le> e\<close> \<open>d < b\<close> \<delta>_def | 
| 31364 | 410 | show False by (cases "t < b") auto | 
| 411 | qed | |
| 412 | ||
| 413 | let ?D1 = "take N D" | |
| 414 | let ?D2 = "drop N D" | |
| 63040 | 415 | define D1 where "D1 = take N D @ [(d, t, b)]" | 
| 416 | define D2 where "D2 = (if b = e then [] else [(b, t, e)]) @ drop (Suc N) D" | |
| 31364 | 417 | |
| 61343 | 418 | from hd_drop_conv_nth[OF \<open>N < length D\<close>] | 
| 419 | have "fst (hd ?D2) = d" using \<open>D ! N = (d, t, e)\<close> by auto | |
| 31364 | 420 | with fine_append_split[OF _ _ append_take_drop_id[symmetric]] | 
| 421 | have fine1: "fine \<delta> (a,d) ?D1" and fine2: "fine \<delta> (d,c) ?D2" | |
| 61343 | 422 | using \<open>N < length D\<close> fine by auto | 
| 31364 | 423 | |
| 424 | have "fine \<delta>1 (a,b) D1" unfolding D1_def | |
| 425 | proof (rule fine_append) | |
| 426 | show "fine \<delta>1 (a, d) ?D1" | |
| 427 | proof (rule fine1[THEN fine_\<delta>_expand]) | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31366diff
changeset | 428 | fix x assume "a \<le> x" "x \<le> d" | 
| 61343 | 429 | hence "x \<le> b" using \<open>d < b\<close> \<open>x \<le> d\<close> by auto | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31366diff
changeset | 430 | thus "\<delta> x \<le> \<delta>1 x" unfolding \<delta>_def by auto | 
| 31364 | 431 | qed | 
| 432 | ||
| 433 | have "b - d < \<delta>1 t" | |
| 61343 | 434 | using mem_fine3[OF fine in_D] \<delta>_def \<open>b \<le> e\<close> \<open>t = b\<close> by auto | 
| 435 | from \<open>d < b\<close> \<open>d \<le> t\<close> \<open>t = b\<close> this | |
| 31364 | 436 | show "fine \<delta>1 (d, b) [(d, t, b)]" using fine_single by auto | 
| 437 | qed | |
| 438 | note rsum1 = I1[OF this] | |
| 439 | ||
| 440 | have drop_split: "drop N D = [D ! N] @ drop (Suc N) D" | |
| 61343 | 441 | using Cons_nth_drop_Suc[OF \<open>N < length D\<close>] by simp | 
| 31364 | 442 | |
| 443 | have fine2: "fine \<delta>2 (e,c) (drop (Suc N) D)" | |
| 444 | proof (cases "drop (Suc N) D = []") | |
| 445 | case True | |
| 446 | note * = fine2[simplified drop_split True D_eq append_Nil2] | |
| 447 | have "e = c" using fine_single_boundaries[OF * refl] by auto | |
| 448 | thus ?thesis unfolding True using fine_Nil by auto | |
| 449 | next | |
| 450 | case False | |
| 451 | note * = fine_append_split[OF fine2 False drop_split] | |
| 452 | from fine_single_boundaries[OF *(1)] | |
| 453 | have "fst (hd (drop (Suc N) D)) = e" using D_eq by auto | |
| 454 | with *(2) have "fine \<delta> (e,c) (drop (Suc N) D)" by auto | |
| 455 | thus ?thesis | |
| 456 | proof (rule fine_\<delta>_expand) | |
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31366diff
changeset | 457 | fix x assume "e \<le> x" and "x \<le> c" | 
| 61343 | 458 | thus "\<delta> x \<le> \<delta>2 x" using \<open>b \<le> e\<close> unfolding \<delta>_def by auto | 
| 31364 | 459 | qed | 
| 460 | qed | |
| 461 | ||
| 462 | have "fine \<delta>2 (b, c) D2" | |
| 463 | proof (cases "e = b") | |
| 464 | case True thus ?thesis using fine2 by (simp add: D1_def D2_def) | |
| 465 | next | |
| 466 | case False | |
| 467 | have "e - b < \<delta>2 b" | |
| 61343 | 468 | using mem_fine3[OF fine in_D] \<delta>_def \<open>d < b\<close> \<open>t = b\<close> by auto | 
| 469 | with False \<open>t = b\<close> \<open>b \<le> e\<close> | |
| 31364 | 470 | show ?thesis using D2_def | 
| 32960 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31366diff
changeset | 471 | by (auto intro!: fine_append[OF _ fine2] fine_single | 
| 
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
 wenzelm parents: 
31366diff
changeset | 472 | simp del: append_Cons) | 
| 31364 | 473 | qed | 
| 474 | note rsum2 = I2[OF this] | |
| 475 | ||
| 476 | have "rsum D f = rsum (take N D) f + rsum [D ! N] f + rsum (drop (Suc N) D) f" | |
| 61343 | 477 | using rsum_append[symmetric] Cons_nth_drop_Suc[OF \<open>N < length D\<close>] by auto | 
| 31364 | 478 | also have "\<dots> = rsum D1 f + rsum D2 f" | 
| 31366 | 479 | by (cases "b = e", auto simp add: D1_def D2_def D_eq rsum_append algebra_simps) | 
| 31364 | 480 | finally have "\<bar>rsum D f - (x1 + x2)\<bar> < \<epsilon>" | 
| 481 | using add_strict_mono[OF rsum1 rsum2] by simp | |
| 482 | } | |
| 483 |   ultimately show "\<exists> \<delta>. gauge {a .. c} \<delta> \<and>
 | |
| 484 | (\<forall>D. fine \<delta> (a,c) D \<longrightarrow> \<bar>rsum D f - (x1 + x2)\<bar> < \<epsilon>)" | |
| 485 | by blast | |
| 486 | next | |
| 487 | case False | |
| 61343 | 488 | hence "a = b \<or> b = c" using \<open>a \<le> b\<close> and \<open>b \<le> c\<close> by auto | 
| 31364 | 489 | thus ?thesis | 
| 490 | proof (rule disjE) | |
| 491 | assume "a = b" hence "x1 = 0" | |
| 61343 | 492 | using \<open>Integral (a, b) f x1\<close> by simp | 
| 493 | thus ?thesis using \<open>a = b\<close> \<open>Integral (b, c) f x2\<close> by simp | |
| 31364 | 494 | next | 
| 495 | assume "b = c" hence "x2 = 0" | |
| 61343 | 496 | using \<open>Integral (b, c) f x2\<close> by simp | 
| 497 | thus ?thesis using \<open>b = c\<close> \<open>Integral (a, b) f x1\<close> by simp | |
| 31364 | 498 | qed | 
| 499 | qed | |
| 31259 
c1b981b71dba
encode gauge-fine partitions with lists instead of functions; remove lots of unnecessary lemmas
 huffman parents: 
31257diff
changeset | 500 | |
| 61343 | 501 | text\<open>Fundamental theorem of calculus (Part I)\<close> | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 502 | |
| 61343 | 503 | text\<open>"Straddle Lemma" : Swartz and Thompson: AMM 95(7) 1988\<close> | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 504 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 505 | lemma strad1: | 
| 53755 | 506 | fixes z x s e :: real | 
| 507 | assumes P: "(\<And>z. z \<noteq> x \<Longrightarrow> \<bar>z - x\<bar> < s \<Longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < e / 2)" | |
| 508 | assumes "\<bar>z - x\<bar> < s" | |
| 509 | shows "\<bar>f z - f x - f' x * (z - x)\<bar> \<le> e / 2 * \<bar>z - x\<bar>" | |
| 510 | proof (cases "z = x") | |
| 511 | case True then show ?thesis by simp | |
| 512 | next | |
| 513 | case False | |
| 514 | then have "inverse (z - x) * (f z - f x - f' x * (z - x)) = (f z - f x) / (z - x) - f' x" | |
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56541diff
changeset | 515 | apply (subst mult.commute) | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53755diff
changeset | 516 | apply (simp add: left_diff_distrib) | 
| 57512 
cc97b347b301
reduced name variants for assoc and commute on plus and mult
 haftmann parents: 
56541diff
changeset | 517 | apply (simp add: mult.assoc divide_inverse) | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53755diff
changeset | 518 | apply (simp add: ring_distribs) | 
| 53755 | 519 | done | 
| 61343 | 520 | moreover from False \<open>\<bar>z - x\<bar> < s\<close> have "\<bar>(f z - f x) / (z - x) - f' x\<bar> < e / 2" | 
| 53755 | 521 | by (rule P) | 
| 522 | ultimately have "\<bar>inverse (z - x)\<bar> * (\<bar>f z - f x - f' x * (z - x)\<bar> * 2) | |
| 523 | \<le> \<bar>inverse (z - x)\<bar> * (e * \<bar>z - x\<bar>)" | |
| 524 | using False by (simp del: abs_inverse add: abs_mult [symmetric] ac_simps) | |
| 525 | with False have "\<bar>f z - f x - f' x * (z - x)\<bar> * 2 \<le> e * \<bar>z - x\<bar>" | |
| 526 | by simp | |
| 527 | then show ?thesis by simp | |
| 528 | qed | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 529 | |
| 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 530 | lemma lemma_straddle: | 
| 31252 | 531 | assumes f': "\<forall>x. a \<le> x & x \<le> b --> DERIV f x :> f'(x)" and "0 < e" | 
| 31253 | 532 |   shows "\<exists>g. gauge {a..b} g &
 | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 533 | (\<forall>x u v. a \<le> u & u \<le> x & x \<le> v & v \<le> b & (v - u) < g(x) | 
| 15094 
a7d1a3fdc30d
conversion of Hyperreal/{Fact,Filter} to Isar scripts
 paulson parents: 
15093diff
changeset | 534 | --> \<bar>(f(v) - f(u)) - (f'(x) * (v - u))\<bar> \<le> e * (v - u))" | 
| 31252 | 535 | proof - | 
| 31253 | 536 |   have "\<forall>x\<in>{a..b}.
 | 
| 15360 | 537 | (\<exists>d > 0. \<forall>u v. u \<le> x & x \<le> v & (v - u) < d --> | 
| 31252 | 538 | \<bar>(f(v) - f(u)) - (f'(x) * (v - u))\<bar> \<le> e * (v - u))" | 
| 31253 | 539 | proof (clarsimp) | 
| 31252 | 540 | fix x :: real assume "a \<le> x" and "x \<le> b" | 
| 541 | with f' have "DERIV f x :> f'(x)" by simp | |
| 542 | then have "\<forall>r>0. \<exists>s>0. \<forall>z. z \<noteq> x \<and> \<bar>z - x\<bar> < s \<longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < r" | |
| 68634 | 543 | by (simp add: has_field_derivative_iff LIM_eq) | 
| 61343 | 544 | with \<open>0 < e\<close> obtain s | 
| 63060 | 545 | where "z \<noteq> x \<Longrightarrow> \<bar>z - x\<bar> < s \<Longrightarrow> \<bar>(f z - f x) / (z - x) - f' x\<bar> < e/2" and "0 < s" for z | 
| 31252 | 546 | by (drule_tac x="e/2" in spec, auto) | 
| 53755 | 547 | with strad1 [of x s f f' e] have strad: | 
| 548 | "\<And>z. \<bar>z - x\<bar> < s \<Longrightarrow> \<bar>f z - f x - f' x * (z - x)\<bar> \<le> e/2 * \<bar>z - x\<bar>" | |
| 549 | by auto | |
| 31252 | 550 | show "\<exists>d>0. \<forall>u v. u \<le> x \<and> x \<le> v \<and> v - u < d \<longrightarrow> \<bar>f v - f u - f' x * (v - u)\<bar> \<le> e * (v - u)" | 
| 551 | proof (safe intro!: exI) | |
| 552 | show "0 < s" by fact | |
| 553 | next | |
| 554 | fix u v :: real assume "u \<le> x" and "x \<le> v" and "v - u < s" | |
| 555 | have "\<bar>f v - f u - f' x * (v - u)\<bar> = | |
| 556 | \<bar>(f v - f x - f' x * (v - x)) + (f x - f u - f' x * (x - u))\<bar>" | |
| 557 | by (simp add: right_diff_distrib) | |
| 558 | also have "\<dots> \<le> \<bar>f v - f x - f' x * (v - x)\<bar> + \<bar>f x - f u - f' x * (x - u)\<bar>" | |
| 559 | by (rule abs_triangle_ineq) | |
| 560 | also have "\<dots> = \<bar>f v - f x - f' x * (v - x)\<bar> + \<bar>f u - f x - f' x * (u - x)\<bar>" | |
| 561 | by (simp add: right_diff_distrib) | |
| 562 | also have "\<dots> \<le> (e/2) * \<bar>v - x\<bar> + (e/2) * \<bar>u - x\<bar>" | |
| 61343 | 563 | using \<open>u \<le> x\<close> \<open>x \<le> v\<close> \<open>v - u < s\<close> by (intro add_mono strad, simp_all) | 
| 31252 | 564 | also have "\<dots> \<le> e * (v - u) / 2 + e * (v - u) / 2" | 
| 61343 | 565 | using \<open>u \<le> x\<close> \<open>x \<le> v\<close> \<open>0 < e\<close> by (intro add_mono, simp_all) | 
| 31252 | 566 | also have "\<dots> = e * (v - u)" | 
| 567 | by simp | |
| 568 | finally show "\<bar>f v - f u - f' x * (v - u)\<bar> \<le> e * (v - u)" . | |
| 569 | qed | |
| 570 | qed | |
| 571 | thus ?thesis | |
| 31253 | 572 | by (simp add: gauge_def) (drule bchoice, auto) | 
| 31252 | 573 | qed | 
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 574 | |
| 35328 | 575 | lemma fundamental_theorem_of_calculus: | 
| 35441 | 576 | assumes "a \<le> b" | 
| 577 | assumes f': "\<forall>x. a \<le> x \<and> x \<le> b \<longrightarrow> DERIV f x :> f'(x)" | |
| 578 | shows "Integral (a, b) f' (f(b) - f(a))" | |
| 579 | proof (cases "a = b") | |
| 580 | assume "a = b" thus ?thesis by simp | |
| 581 | next | |
| 61343 | 582 | assume "a \<noteq> b" with \<open>a \<le> b\<close> have "a < b" by simp | 
| 35441 | 583 | show ?thesis | 
| 584 | proof (simp add: Integral_def2, clarify) | |
| 585 | fix e :: real assume "0 < e" | |
| 61343 | 586 | with \<open>a < b\<close> have "0 < e / (b - a)" by simp | 
| 35441 | 587 | |
| 588 | from lemma_straddle [OF f' this] | |
| 589 |     obtain \<delta> where "gauge {a..b} \<delta>"
 | |
| 63060 | 590 | and \<delta>: "\<lbrakk>a \<le> u; u \<le> x; x \<le> v; v \<le> b; v - u < \<delta> x\<rbrakk> \<Longrightarrow> | 
| 591 | \<bar>f v - f u - f' x * (v - u)\<bar> \<le> e * (v - u) / (b - a)" for x u v by auto | |
| 35441 | 592 | |
| 593 | have "\<forall>D. fine \<delta> (a, b) D \<longrightarrow> \<bar>rsum D f' - (f b - f a)\<bar> \<le> e" | |
| 594 | proof (clarify) | |
| 595 | fix D assume D: "fine \<delta> (a, b) D" | |
| 596 | hence "(\<Sum>(u, x, v)\<leftarrow>D. f v - f u) = f b - f a" | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63627diff
changeset | 597 | by (rule fine_sum_list_eq_diff) | 
| 35441 | 598 | hence "\<bar>rsum D f' - (f b - f a)\<bar> = \<bar>rsum D f' - (\<Sum>(u, x, v)\<leftarrow>D. f v - f u)\<bar>" | 
| 599 | by simp | |
| 600 | also have "\<dots> = \<bar>(\<Sum>(u, x, v)\<leftarrow>D. f v - f u) - rsum D f'\<bar>" | |
| 601 | by (rule abs_minus_commute) | |
| 602 | also have "\<dots> = \<bar>\<Sum>(u, x, v)\<leftarrow>D. (f v - f u) - f' x * (v - u)\<bar>" | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63627diff
changeset | 603 | by (simp only: rsum_def sum_list_subtractf split_def) | 
| 35441 | 604 | also have "\<dots> \<le> (\<Sum>(u, x, v)\<leftarrow>D. \<bar>(f v - f u) - f' x * (v - u)\<bar>)" | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63627diff
changeset | 605 | by (rule ord_le_eq_trans [OF sum_list_abs], simp add: o_def split_def) | 
| 35441 | 606 | also have "\<dots> \<le> (\<Sum>(u, x, v)\<leftarrow>D. (e / (b - a)) * (v - u))" | 
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63627diff
changeset | 607 | apply (rule sum_list_mono, clarify, rename_tac u x v) | 
| 35441 | 608 | using D apply (simp add: \<delta> mem_fine mem_fine2 mem_fine3) | 
| 609 | done | |
| 610 | also have "\<dots> = e" | |
| 63882 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63627diff
changeset | 611 | using fine_sum_list_eq_diff [OF D, where f="\<lambda>x. x"] | 
| 
018998c00003
renamed listsum -> sum_list, listprod ~> prod_list
 nipkow parents: 
63627diff
changeset | 612 | unfolding split_def sum_list_const_mult | 
| 61343 | 613 | using \<open>a < b\<close> by simp | 
| 35441 | 614 | finally show "\<bar>rsum D f' - (f b - f a)\<bar> \<le> e" . | 
| 615 | qed | |
| 616 | ||
| 61343 | 617 |     with \<open>gauge {a..b} \<delta>\<close>
 | 
| 35441 | 618 |     show "\<exists>\<delta>. gauge {a..b} \<delta> \<and> (\<forall>D. fine \<delta> (a, b) D \<longrightarrow> \<bar>rsum D f' - (f b - f a)\<bar> \<le> e)"
 | 
| 619 | by auto | |
| 620 | qed | |
| 621 | qed | |
| 13958 | 622 | |
| 15093 
49ede01e9ee6
conversion of Integration and NSPrimes to Isar scripts
 paulson parents: 
13958diff
changeset | 623 | end |