author | paulson |
Thu, 01 Jan 2004 10:06:32 +0100 | |
changeset 14334 | 6137d24eef79 |
parent 14329 | ff3210fe968f |
child 14336 | 8f731d3cd65b |
permissions | -rw-r--r-- |
13958 | 1 |
(* Title : Integration.ML |
2 |
Author : Jacques D. Fleuriot |
|
3 |
Copyright : 2000 University of Edinburgh |
|
4 |
Description : Theory of integration (c.f. Harison's HOL development) |
|
5 |
*) |
|
6 |
||
7 |
Goalw [psize_def] "a = b ==> psize (%n. if n = 0 then a else b) = 0"; |
|
8 |
by Auto_tac; |
|
9 |
qed "partition_zero"; |
|
10 |
Addsimps [partition_zero]; |
|
11 |
||
12 |
Goalw [psize_def] "a < b ==> psize (%n. if n = 0 then a else b) = 1"; |
|
13 |
by Auto_tac; |
|
14 |
by (rtac some_equality 1); |
|
15 |
by Auto_tac; |
|
16 |
by (dres_inst_tac [("x","1")] spec 1); |
|
17 |
by Auto_tac; |
|
18 |
qed "partition_one"; |
|
19 |
Addsimps [partition_one]; |
|
20 |
||
21 |
Goalw [partition_def] |
|
22 |
"a <= b ==> partition(a,b)(%n. if n = 0 then a else b)"; |
|
23 |
by (auto_tac (claset(),simpset() addsimps [real_le_less])); |
|
24 |
qed "partition_single"; |
|
25 |
Addsimps [partition_single]; |
|
26 |
||
27 |
Goalw [partition_def] "partition(a,b) D ==> (D(0) = a)"; |
|
28 |
by Auto_tac; |
|
29 |
qed "partition_lhs"; |
|
30 |
||
31 |
Goalw [partition_def] |
|
32 |
"(partition(a,b) D) = \ |
|
33 |
\ ((D 0 = a) & \ |
|
34 |
\ (ALL n. n < (psize D) --> D n < D(Suc n)) & \ |
|
35 |
\ (ALL n. (psize D) <= n --> (D n = b)))"; |
|
36 |
by Auto_tac; |
|
37 |
by (ALLGOALS(subgoal_tac "psize D = N")); |
|
38 |
by Auto_tac; |
|
39 |
by (ALLGOALS(simp_tac (simpset() addsimps [psize_def]))); |
|
40 |
by (ALLGOALS(rtac some_equality)); |
|
41 |
by (Blast_tac 1 THEN Blast_tac 2); |
|
42 |
by (ALLGOALS(rtac ccontr)); |
|
43 |
by (ALLGOALS(dtac (ARITH_PROVE "n ~= (N::nat) ==> n < N | N < n"))); |
|
44 |
by (Step_tac 1); |
|
45 |
by (dres_inst_tac [("x","Na")] spec 1); |
|
46 |
by (rotate_tac 3 1); |
|
47 |
by (dres_inst_tac [("x","Suc Na")] spec 1); |
|
48 |
by (Asm_full_simp_tac 1); |
|
49 |
by (rotate_tac 2 1); |
|
50 |
by (dres_inst_tac [("x","N")] spec 1); |
|
51 |
by (Asm_full_simp_tac 1); |
|
52 |
by (dres_inst_tac [("x","Na")] spec 1); |
|
53 |
by (rotate_tac 3 1); |
|
54 |
by (dres_inst_tac [("x","Suc Na")] spec 1); |
|
55 |
by (Asm_full_simp_tac 1); |
|
56 |
by (rotate_tac 2 1); |
|
57 |
by (dres_inst_tac [("x","N")] spec 1); |
|
58 |
by (Asm_full_simp_tac 1); |
|
59 |
qed "partition"; |
|
60 |
||
61 |
Goal "partition(a,b) D ==> (D(psize D) = b)"; |
|
62 |
by (dtac (partition RS subst) 1); |
|
63 |
by (Step_tac 1); |
|
64 |
by (REPEAT(dres_inst_tac [("x","psize D")] spec 1)); |
|
65 |
by Auto_tac; |
|
66 |
qed "partition_rhs"; |
|
67 |
||
68 |
Goal "[|partition(a,b) D; psize D <= n |] ==> (D n = b)"; |
|
69 |
by (dtac (partition RS subst) 1); |
|
70 |
by (Blast_tac 1); |
|
71 |
qed "partition_rhs2"; |
|
72 |
||
73 |
Goal |
|
74 |
"partition(a,b) D & m + Suc d <= n & n <= (psize D) --> D(m) < D(m + Suc d)"; |
|
75 |
by (induct_tac "d" 1); |
|
76 |
by Auto_tac; |
|
77 |
by (ALLGOALS(dtac (partition RS subst))); |
|
78 |
by (Step_tac 1); |
|
79 |
by (ALLGOALS(dtac (ARITH_PROVE "Suc m <= n ==> m < n"))); |
|
80 |
by (ALLGOALS(dtac less_le_trans)); |
|
81 |
by (assume_tac 1 THEN assume_tac 2); |
|
82 |
by (ALLGOALS(blast_tac (claset() addIs [real_less_trans]))); |
|
83 |
qed_spec_mp "lemma_partition_lt_gen"; |
|
84 |
||
85 |
Goal "m < n ==> EX d. n = m + Suc d"; |
|
86 |
by (auto_tac (claset(),simpset() addsimps [less_iff_Suc_add])); |
|
87 |
qed "less_eq_add_Suc"; |
|
88 |
||
89 |
Goal "[|partition(a,b) D; m < n; n <= (psize D)|] ==> D(m) < D(n)"; |
|
90 |
by (auto_tac (claset() addDs [less_eq_add_Suc] addIs [lemma_partition_lt_gen], |
|
91 |
simpset())); |
|
92 |
qed "partition_lt_gen"; |
|
93 |
||
94 |
Goal "partition(a,b) D ==> (n < (psize D) --> D(0) < D(Suc n))"; |
|
95 |
by (dtac (partition RS subst) 1); |
|
96 |
by (induct_tac "n" 1); |
|
97 |
by (Blast_tac 1); |
|
98 |
by (blast_tac (claset() addSDs [leI] addDs |
|
99 |
[(ARITH_PROVE "m <= n ==> m <= Suc n"), |
|
100 |
le_less_trans,real_less_trans]) 1); |
|
101 |
qed_spec_mp "partition_lt"; |
|
102 |
||
103 |
Goal "partition(a,b) D ==> a <= b"; |
|
104 |
by (ftac (partition RS subst) 1); |
|
105 |
by (Step_tac 1); |
|
106 |
by (rotate_tac 2 1); |
|
107 |
by (dres_inst_tac [("x","psize D")] spec 1); |
|
108 |
by (Step_tac 1); |
|
109 |
by (rtac ccontr 1); |
|
110 |
by (case_tac "psize D = 0" 1); |
|
111 |
by (Step_tac 1); |
|
112 |
by (dres_inst_tac [("n","psize D - 1")] partition_lt 2); |
|
113 |
by Auto_tac; |
|
114 |
qed "partition_le"; |
|
115 |
||
116 |
Goal "[|partition(a,b) D; n < (psize D)|] ==> D(n) < D(psize D)"; |
|
117 |
by (auto_tac (claset() addIs [partition_lt_gen],simpset())); |
|
118 |
qed "partition_gt"; |
|
119 |
||
120 |
Goal "partition(a,b) D ==> ((a = b) = (psize D = 0))"; |
|
121 |
by (ftac (partition RS subst) 1); |
|
122 |
by (Step_tac 1); |
|
123 |
by (rotate_tac 2 1); |
|
124 |
by (dres_inst_tac [("x","psize D")] spec 1); |
|
125 |
by (rtac ccontr 1); |
|
126 |
by (dres_inst_tac [("n","psize D - 1")] partition_lt 1); |
|
127 |
by (Blast_tac 3 THEN Auto_tac); |
|
128 |
qed "partition_eq"; |
|
129 |
||
130 |
Goal "partition(a,b) D ==> a <= D(r)"; |
|
131 |
by (ftac (partition RS subst) 1); |
|
132 |
by (Step_tac 1); |
|
133 |
by (induct_tac "r" 1); |
|
134 |
by (cut_inst_tac [("m","Suc n"),("n","psize D")] |
|
135 |
(ARITH_PROVE "m < n | n <= (m::nat)") 2); |
|
136 |
by (Step_tac 1); |
|
137 |
by (eres_inst_tac [("j","D n")] real_le_trans 1); |
|
138 |
by (dres_inst_tac [("x","n")] spec 1); |
|
139 |
by (blast_tac (claset() addIs [less_trans,order_less_imp_le]) 1); |
|
140 |
by (res_inst_tac [("j","b")] real_le_trans 1); |
|
141 |
by (etac partition_le 1); |
|
142 |
by (Blast_tac 1); |
|
143 |
qed "partition_lb"; |
|
144 |
||
145 |
Goal "[| partition(a,b) D; psize D ~= 0 |] ==> a < D(Suc n)"; |
|
146 |
by (res_inst_tac [("t","a")] (partition_lhs RS subst) 1); |
|
147 |
by (assume_tac 1); |
|
148 |
by (cut_inst_tac [("m","psize D"),("n","Suc n")] |
|
149 |
(ARITH_PROVE "m < n | n <= (m::nat)") 1); |
|
150 |
by (ftac (partition RS subst) 1); |
|
151 |
by (Step_tac 1); |
|
152 |
by (rotate_tac 3 1); |
|
153 |
by (dres_inst_tac [("x","Suc n")] spec 1); |
|
154 |
by (etac impE 1); |
|
155 |
by (etac less_imp_le 1); |
|
156 |
by (ftac partition_rhs 1); |
|
157 |
by (dtac partition_gt 1 THEN assume_tac 1); |
|
158 |
by (Asm_simp_tac 1); |
|
159 |
by (blast_tac (claset() addIs [partition_lt,less_le_trans]) 1); |
|
160 |
qed "partition_lb_lt"; |
|
161 |
||
162 |
Goal "partition(a,b) D ==> D(r) <= b"; |
|
163 |
by (ftac (partition RS subst) 1); |
|
164 |
by (cut_inst_tac [("m","r"),("n","psize D")] |
|
165 |
(ARITH_PROVE "m < n | n <= (m::nat)") 1); |
|
166 |
by (Step_tac 1); |
|
167 |
by (Blast_tac 2); |
|
168 |
by (subgoal_tac "ALL x. D((psize D) - x) <= b" 1); |
|
169 |
by (rotate_tac 4 1); |
|
170 |
by (dres_inst_tac [("x","psize D - r")] spec 1); |
|
171 |
by (dtac (ARITH_PROVE "p < m ==> m - (m - p) = (p::nat)") 1); |
|
172 |
by (thin_tac "ALL n. psize D <= n --> D n = b" 1); |
|
173 |
by (Asm_full_simp_tac 1); |
|
174 |
by (Step_tac 1); |
|
175 |
by (induct_tac "x" 1); |
|
176 |
by (Simp_tac 1 THEN Blast_tac 1); |
|
177 |
by (case_tac "psize D - Suc n = 0" 1); |
|
178 |
by (thin_tac "ALL n. psize D <= n --> D n = b" 1); |
|
179 |
by (asm_simp_tac (simpset() addsimps [partition_le]) 1); |
|
180 |
by (rtac real_le_trans 1 THEN assume_tac 2); |
|
181 |
by (rtac (ARITH_PROVE "0 < m - Suc n ==> (m - n) = Suc(m - Suc n)" |
|
182 |
RS ssubst) 1); |
|
183 |
by (dres_inst_tac [("x","psize D - Suc n")] spec 2); |
|
184 |
by (thin_tac "ALL n. psize D <= n --> D n = b" 2); |
|
185 |
by (Asm_full_simp_tac 2); |
|
14329 | 186 |
by (arith_tac 1); |
13958 | 187 |
qed "partition_ub"; |
188 |
||
189 |
Goal "[| partition(a,b) D; n < psize D |] ==> D(n) < b"; |
|
190 |
by (blast_tac (claset() addIs [partition_rhs RS subst] addIs |
|
191 |
[partition_gt]) 1); |
|
192 |
qed "partition_ub_lt"; |
|
193 |
||
194 |
Goal "[| partition (a, b) D1; partition (b, c) D2 |] \ |
|
195 |
\ ==> (ALL n. \ |
|
196 |
\ n < psize D1 + psize D2 --> \ |
|
197 |
\ (if n < psize D1 then D1 n else D2 (n - psize D1)) \ |
|
198 |
\ < (if Suc n < psize D1 then D1 (Suc n) \ |
|
199 |
\ else D2 (Suc n - psize D1))) & \ |
|
200 |
\ (ALL n. \ |
|
201 |
\ psize D1 + psize D2 <= n --> \ |
|
202 |
\ (if n < psize D1 then D1 n else D2 (n - psize D1)) = \ |
|
203 |
\ (if psize D1 + psize D2 < psize D1 then D1 (psize D1 + psize D2) \ |
|
204 |
\ else D2 (psize D1 + psize D2 - psize D1)))"; |
|
205 |
by (Step_tac 1); |
|
206 |
by (auto_tac (claset() addIs [partition_lt_gen],simpset())); |
|
207 |
by (dres_inst_tac [("m","psize D1")] |
|
208 |
(leI RSN(2,ARITH_PROVE "[| n < m; m <= Suc n |] ==> m = Suc n")) 1); |
|
209 |
by (assume_tac 1); |
|
210 |
by (auto_tac (claset() addSIs [partition_lt_gen], |
|
211 |
simpset() addsimps [partition_lhs,partition_ub_lt])); |
|
212 |
by (blast_tac (claset() addIs [ARITH_PROVE "~n < m ==> n - m < Suc n - m"]) 1); |
|
213 |
by (auto_tac (claset(),simpset() addsimps [ |
|
214 |
ARITH_PROVE "~ n < m ==> (Suc n - m <= p) = (Suc n <= m + p)"])); |
|
215 |
by (dtac (ARITH_PROVE "p + q <= (n::nat) ==> q <= n - p") 1); |
|
216 |
by (auto_tac (claset() addSIs [partition_rhs2], simpset() addsimps |
|
217 |
[partition_rhs])); |
|
218 |
qed "lemma_partition_append1"; |
|
219 |
||
220 |
Goal "[| partition (a, b) D1; partition (b, c) D2; N < psize D1 |] \ |
|
221 |
\ ==> D1(N) < D2 (psize D2)"; |
|
222 |
by (res_inst_tac [("y","D1(psize D1)")] order_less_le_trans 1); |
|
223 |
by (etac partition_gt 1 THEN assume_tac 1); |
|
224 |
by (auto_tac (claset(),simpset() addsimps [partition_rhs,partition_le])); |
|
225 |
qed "lemma_psize1"; |
|
226 |
||
227 |
Goal "[| partition (a, b) D1; partition (b, c) D2 |] \ |
|
228 |
\ ==> psize (%n. if n < psize D1 then D1 n else D2 (n - psize D1)) = \ |
|
229 |
\ psize D1 + psize D2"; |
|
230 |
by (res_inst_tac |
|
231 |
[("D2","%n. if n < psize D1 then D1 n else D2 (n - psize D1)")] |
|
232 |
((psize_def RS meta_eq_to_obj_eq) RS ssubst) 1); |
|
233 |
by (rtac some1_equality 1); |
|
234 |
by (blast_tac (claset() addSIs [lemma_partition_append1]) 2); |
|
235 |
by (rtac ex1I 1 THEN rtac lemma_partition_append1 1); |
|
236 |
by Auto_tac; |
|
237 |
by (rtac ccontr 1); |
|
238 |
by (dtac (ARITH_PROVE "m ~= n ==> m < n | n < (m::nat)") 1); |
|
239 |
by (Step_tac 1); |
|
240 |
by (rotate_tac 3 1); |
|
241 |
by (dres_inst_tac [("x","psize D1 + psize D2")] spec 1); |
|
242 |
by Auto_tac; |
|
243 |
by (case_tac "N < psize D1" 1); |
|
244 |
by (auto_tac (claset() addDs [lemma_psize1],simpset())); |
|
245 |
by (dtac leI 1); |
|
246 |
by (dtac (ARITH_PROVE "[|p <= n; n < p + m|] ==> n - p < (m::nat)") 1); |
|
247 |
by (assume_tac 1); |
|
248 |
by (dres_inst_tac [("a","b"),("b","c")] partition_gt 1); |
|
249 |
by Auto_tac; |
|
250 |
by (dres_inst_tac [("x","psize D1 + psize D2")] spec 1); |
|
251 |
by (auto_tac (claset(),simpset() addsimps [partition_rhs2])); |
|
252 |
qed "lemma_partition_append2"; |
|
253 |
||
254 |
Goalw [tpart_def] |
|
255 |
"[|psize D = 0; tpart(a,b) (D,p)|] ==> a = b"; |
|
256 |
by (auto_tac (claset(),simpset() addsimps [partition_eq])); |
|
257 |
qed "tpart_eq_lhs_rhs"; |
|
258 |
||
259 |
Goalw [tpart_def] "tpart(a,b) (D,p) ==> partition(a,b) D"; |
|
260 |
by Auto_tac; |
|
261 |
qed "tpart_partition"; |
|
262 |
||
263 |
Goal "[| tpart(a,b) (D1,p1); fine(g) (D1,p1); \ |
|
264 |
\ tpart(b,c) (D2,p2); fine(g) (D2,p2) |] \ |
|
265 |
\ ==> EX D p. tpart(a,c) (D,p) & fine(g) (D,p)"; |
|
266 |
by (res_inst_tac |
|
267 |
[("x","%n. if n < psize D1 then D1 n else D2 (n - psize D1)")] exI 1); |
|
268 |
by (res_inst_tac |
|
269 |
[("x","%n. if n < psize D1 then p1 n else p2 (n - psize D1)")] exI 1); |
|
270 |
by (cut_inst_tac [("m","psize D1")] (ARITH_PROVE "0 < (m::nat) | m = 0") 1); |
|
271 |
by (auto_tac (claset() addDs [tpart_eq_lhs_rhs],simpset())); |
|
272 |
by (asm_full_simp_tac (simpset() addsimps [fine_def, |
|
273 |
[tpart_partition,tpart_partition] MRS lemma_partition_append2]) 2); |
|
274 |
by Auto_tac; |
|
275 |
by (dres_inst_tac [("m","psize D1"),("n","n")] |
|
276 |
(leI RSN(2,ARITH_PROVE "[| n < m; m <= Suc n |] ==> m = Suc n")) 2); |
|
277 |
by Auto_tac; |
|
278 |
by (dtac (tpart_partition RS partition_rhs) 2); |
|
279 |
by (dtac (tpart_partition RS partition_lhs) 2); |
|
280 |
by Auto_tac; |
|
281 |
by (rotate_tac 3 2); |
|
282 |
by (dres_inst_tac [("x","n - psize D1")] spec 2); |
|
283 |
by (auto_tac (claset(),simpset() addsimps |
|
284 |
[ARITH_PROVE "[| (0::nat) < p; p <= n |] ==> (n - p < m) = (n < p + m)", |
|
285 |
ARITH_PROVE "~n < p ==> Suc n - p = Suc(n - p)"])); |
|
286 |
by (auto_tac (claset(),simpset() addsimps [tpart_def, |
|
287 |
ARITH_PROVE "~n < p ==> Suc n - p = Suc(n - p)"])); |
|
288 |
by (dres_inst_tac [("m","psize D1"),("n","n")] |
|
289 |
(leI RSN(2,ARITH_PROVE "[| n < m; m <= Suc n |] ==> m = Suc n")) 2); |
|
290 |
by Auto_tac; |
|
291 |
by (dtac partition_rhs 2); |
|
292 |
by (dtac partition_lhs 2); |
|
293 |
by Auto_tac; |
|
294 |
by (rtac (partition RS ssubst) 1); |
|
295 |
by (rtac (lemma_partition_append2 RS ssubst) 1); |
|
296 |
by (rtac conjI 3); |
|
297 |
by (dtac lemma_partition_append1 4); |
|
298 |
by (auto_tac (claset(),simpset() addsimps [partition_lhs,partition_rhs])); |
|
299 |
qed "partition_append"; |
|
300 |
||
301 |
(* ------------------------------------------------------------------------ *) |
|
302 |
(* We can always find a division which is fine wrt any gauge *) |
|
303 |
(* ------------------------------------------------------------------------ *) |
|
304 |
||
305 |
Goal "[| a <= b; gauge(%x. a <= x & x <= b) g |] \ |
|
306 |
\ ==> EX D p. tpart(a,b) (D,p) & fine g (D,p)"; |
|
307 |
by (cut_inst_tac |
|
308 |
[("P","%(u,v). a <= u & v <= b --> (EX D p. tpart(u,v) (D,p) & fine(g) (D,p))")] |
|
309 |
lemma_BOLZANO2 1); |
|
310 |
by (Step_tac 1); |
|
311 |
by (REPEAT(blast_tac (claset() addIs [real_le_trans]) 1)); |
|
312 |
by (auto_tac (claset() addIs [partition_append],simpset())); |
|
313 |
by (case_tac "a <= x & x <= b" 1); |
|
314 |
by (res_inst_tac [("x","1")] exI 2); |
|
315 |
by Auto_tac; |
|
316 |
by (res_inst_tac [("x","g x")] exI 1); |
|
317 |
by (auto_tac (claset(),simpset() addsimps [gauge_def])); |
|
318 |
by (res_inst_tac [("x","%n. if n = 0 then aa else ba")] exI 1); |
|
319 |
by (res_inst_tac [("x","%n. if n = 0 then x else ba")] exI 1); |
|
320 |
by (auto_tac (claset(),simpset() addsimps [tpart_def,fine_def])); |
|
321 |
qed "partition_exists"; |
|
322 |
||
323 |
(* ------------------------------------------------------------------------ *) |
|
324 |
(* Lemmas about combining gauges *) |
|
325 |
(* ------------------------------------------------------------------------ *) |
|
326 |
||
327 |
Goalw [gauge_def] |
|
328 |
"[| gauge(E) g1; gauge(E) g2 |] \ |
|
329 |
\ ==> gauge(E) (%x. if g1(x) < g2(x) then g1(x) else g2(x))"; |
|
330 |
by Auto_tac; |
|
331 |
qed "gauge_min"; |
|
332 |
||
333 |
Goalw [fine_def] |
|
334 |
"fine (%x. if g1(x) < g2(x) then g1(x) else g2(x)) (D,p) \ |
|
335 |
\ ==> fine(g1) (D,p) & fine(g2) (D,p)"; |
|
336 |
by (auto_tac (claset(),simpset() addsimps [split_if])); |
|
337 |
qed "fine_min"; |
|
338 |
||
339 |
||
340 |
(* ------------------------------------------------------------------------ *) |
|
341 |
(* The integral is unique if it exists *) |
|
342 |
(* ------------------------------------------------------------------------ *) |
|
343 |
||
344 |
Goalw [Integral_def] |
|
345 |
"[| a <= b; Integral(a,b) f k1; Integral(a,b) f k2 |] ==> k1 = k2"; |
|
346 |
by Auto_tac; |
|
347 |
by (REPEAT(dres_inst_tac [("x","abs(k1 - k2)/2")] spec 1)); |
|
348 |
by (Auto_tac THEN TRYALL(arith_tac)); |
|
349 |
by (dtac gauge_min 1 THEN assume_tac 1); |
|
350 |
by (dres_inst_tac [("g","%x. if g x < ga x then g x else ga x")] |
|
351 |
partition_exists 1 THEN assume_tac 1); |
|
352 |
by Auto_tac; |
|
353 |
by (dtac fine_min 1); |
|
354 |
by (REPEAT(dtac spec 1) THEN Auto_tac); |
|
355 |
by (subgoal_tac |
|
356 |
"abs((rsum(D,p) f - k2) - (rsum(D,p) f - k1)) < abs(k1 - k2)" 1); |
|
357 |
by (arith_tac 1); |
|
14334 | 358 |
by (dtac add_strict_mono 1 THEN assume_tac 1); |
14305
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14294
diff
changeset
|
359 |
by (auto_tac (claset(), |
14334 | 360 |
HOL_ss addsimps [left_distrib RS sym, |
14305
f17ca9f6dc8c
tidying first part of HyperArith0.ML, using generic lemmas
paulson
parents:
14294
diff
changeset
|
361 |
real_mult_2_right RS sym, mult_less_cancel_right])); |
13958 | 362 |
by (ALLGOALS(arith_tac)); |
363 |
qed "Integral_unique"; |
|
364 |
||
365 |
Goalw [Integral_def] "Integral(a,a) f 0"; |
|
366 |
by Auto_tac; |
|
367 |
by (res_inst_tac [("x","%x. 1")] exI 1); |
|
368 |
by (auto_tac (claset() addDs [partition_eq],simpset() addsimps [gauge_def, |
|
369 |
tpart_def,rsum_def])); |
|
370 |
qed "Integral_zero"; |
|
371 |
Addsimps [Integral_zero]; |
|
372 |
||
373 |
Goal "sumr 0 m (%n. D (Suc n) - D n) = D(m) - D 0"; |
|
374 |
by (induct_tac "m" 1); |
|
375 |
by Auto_tac; |
|
376 |
qed "sumr_partition_eq_diff_bounds"; |
|
377 |
Addsimps [sumr_partition_eq_diff_bounds]; |
|
378 |
||
379 |
Goal "a <= b ==> Integral(a,b) (%x. 1) (b - a)"; |
|
380 |
by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac); |
|
381 |
by (auto_tac (claset(),simpset() addsimps [rsum_def,Integral_def])); |
|
382 |
by (res_inst_tac [("x","%x. b - a")] exI 1); |
|
383 |
by (auto_tac (claset(),simpset() addsimps [gauge_def, |
|
384 |
abs_interval_iff,tpart_def,partition])); |
|
385 |
qed "Integral_eq_diff_bounds"; |
|
386 |
||
387 |
Goal "a <= b ==> Integral(a,b) (%x. c) (c*(b - a))"; |
|
388 |
by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac); |
|
389 |
by (auto_tac (claset(),simpset() addsimps [rsum_def,Integral_def])); |
|
390 |
by (res_inst_tac [("x","%x. b - a")] exI 1); |
|
391 |
by (auto_tac (claset(),simpset() addsimps |
|
392 |
[sumr_mult RS sym,gauge_def,abs_interval_iff, |
|
14334 | 393 |
right_diff_distrib RS sym,partition,tpart_def])); |
13958 | 394 |
qed "Integral_mult_const"; |
395 |
||
396 |
Goal "[| a <= b; Integral(a,b) f k |] ==> Integral(a,b) (%x. c * f x) (c * k)"; |
|
397 |
by (dtac real_le_imp_less_or_eq 1); |
|
398 |
by (auto_tac (claset() addDs [[real_le_refl,Integral_zero] MRS Integral_unique], |
|
399 |
simpset())); |
|
400 |
by (auto_tac (claset(),simpset() addsimps [rsum_def,Integral_def, |
|
401 |
sumr_mult RS sym,real_mult_assoc])); |
|
14294
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
paulson
parents:
14277
diff
changeset
|
402 |
by (res_inst_tac [("a2","c")] ((abs_ge_zero RS real_le_imp_less_or_eq) |
13958 | 403 |
RS disjE) 1); |
404 |
by (dtac sym 2); |
|
405 |
by (Asm_full_simp_tac 2 THEN Blast_tac 2); |
|
406 |
by (dres_inst_tac [("x","e/abs c")] spec 1 THEN Auto_tac); |
|
14334 | 407 |
by (asm_full_simp_tac (simpset() addsimps [zero_less_mult_iff, |
13958 | 408 |
real_divide_def]) 1); |
409 |
by (rtac exI 1 THEN Auto_tac); |
|
410 |
by (REPEAT(dtac spec 1) THEN Auto_tac); |
|
411 |
by (res_inst_tac [("z1","inverse(abs c)")] (real_mult_less_iff1 RS iffD1) 1); |
|
412 |
by (fold_tac [real_divide_def]); |
|
14334 | 413 |
by (auto_tac (claset(), |
414 |
simpset() addsimps [right_diff_distrib RS sym, |
|
415 |
abs_mult, real_mult_assoc RS sym, |
|
416 |
ARITH_PROVE "c ~= 0 ==> abs (c::real) ~= 0", |
|
417 |
positive_imp_inverse_positive])); |
|
13958 | 418 |
qed "Integral_mult"; |
419 |
||
420 |
(* ------------------------------------------------------------------------ *) |
|
421 |
(* Fundamental theorem of calculus (Part I) *) |
|
422 |
(* ------------------------------------------------------------------------ *) |
|
423 |
||
424 |
(* "Straddle Lemma" : Swartz & Thompson: AMM 95(7) 1988 *) |
|
425 |
||
426 |
Goal "ALL x. P(x) --> (EX y. Q x y) ==> EX f. (ALL x. P(x) --> Q x (f x))"; |
|
427 |
by (cut_inst_tac [("S","Collect (%x. P x)"),("Q","Q")] (CLAIM_SIMP |
|
428 |
"(ALL x:S. (EX y. Q x y)) --> (EX f. (ALL x:S. Q x (f x)))" [bchoice]) 1); |
|
429 |
by Auto_tac; |
|
430 |
qed "choiceP"; |
|
431 |
||
432 |
Goal "ALL x. P(x) --> (EX y. R(y) & (EX z. Q x y z)) ==> \ |
|
433 |
\ EX f fa. (ALL x. P(x) --> R(f x) & Q x (f x) (fa x))"; |
|
434 |
by (dtac (CLAIM "(ALL x. P(x) --> (EX y. R(y) & (EX z. Q x y z))) = \ |
|
435 |
\ (ALL x. P(x) --> (EX y z. R(y) & Q x y z))" RS iffD1) 1);; |
|
436 |
by (dtac choiceP 1 THEN Step_tac 1); |
|
437 |
by (dtac choiceP 1 THEN Auto_tac); |
|
438 |
qed "choiceP2"; |
|
439 |
||
440 |
Goal "ALL x. (EX y. R(y) & (EX z. Q x y z)) ==> \ |
|
441 |
\ EX f fa. (ALL x. R(f x) & Q x (f x) (fa x))"; |
|
442 |
by (dtac (CLAIM "(ALL x. (EX y. R(y) & (EX z. Q x y z))) = \ |
|
443 |
\ (ALL x. (EX y z. R(y) & Q x y z))" RS iffD1) 1);; |
|
444 |
by (dtac choice 1 THEN Step_tac 1); |
|
445 |
by (dtac choice 1 THEN Auto_tac); |
|
446 |
qed "choice2"; |
|
447 |
||
448 |
(* new simplifications e.g. (y < x/n) = (y * n < x) are a real nuisance |
|
449 |
they break the original proofs and make new proofs longer! *) |
|
450 |
Goalw [gauge_def] |
|
451 |
"[| ALL x. a <= x & x <= b --> DERIV f x :> f'(x); 0 < e |]\ |
|
452 |
\ ==> EX g. gauge(%x. a <= x & x <= b) g & \ |
|
453 |
\ (ALL x u v. a <= u & u <= x & x <= v & v <= b & (v - u) < g(x) \ |
|
454 |
\ --> abs((f(v) - f(u)) - (f'(x) * (v - u))) <= e * (v - u))"; |
|
455 |
by (subgoal_tac "ALL x. a <= x & x <= b --> \ |
|
456 |
\ (EX d. 0 < d & \ |
|
457 |
\ (ALL u v. u <= x & x <= v & (v - u) < d --> \ |
|
458 |
\ abs((f(v) - f(u)) - (f'(x) * (v - u))) <= e * (v - u)))" 1); |
|
459 |
by (dtac choiceP 1); |
|
460 |
by Auto_tac; |
|
461 |
by (dtac spec 1 THEN Auto_tac); |
|
462 |
by (auto_tac (claset(),simpset() addsimps [DERIV_iff2,LIM_def])); |
|
463 |
by (dres_inst_tac [("x","e/2")] spec 1); |
|
464 |
by Auto_tac; |
|
465 |
by (subgoal_tac "ALL z. abs(z - x) < s --> \ |
|
466 |
\ abs((f(z) - f(x)) - (f'(x) * (z - x))) <= (e / 2) * abs(z - x)" 1); |
|
467 |
by Auto_tac; |
|
468 |
by (case_tac "0 < abs(z - x)" 2); |
|
469 |
by (dtac (ARITH_PROVE "~ 0 < abs z ==> z = (0::real)") 3); |
|
470 |
by (asm_full_simp_tac (simpset() addsimps |
|
471 |
[ARITH_PROVE "(z - x = 0) = (z = (x::real))"]) 3); |
|
472 |
by (dres_inst_tac [("x","z")] spec 2); |
|
473 |
by (res_inst_tac [("z1","abs(inverse(z - x))")] |
|
474 |
(real_mult_le_cancel_iff2 RS iffD1) 2); |
|
475 |
by (Asm_full_simp_tac 2); |
|
14294
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
paulson
parents:
14277
diff
changeset
|
476 |
by (asm_full_simp_tac (simpset() |
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
paulson
parents:
14277
diff
changeset
|
477 |
delsimps [abs_inverse] |
f4d806fd72ce
absolute value theorems moved to HOL/Ring_and_Field
paulson
parents:
14277
diff
changeset
|
478 |
addsimps [abs_mult RS sym, real_mult_assoc RS sym]) 2); |
13958 | 479 |
by (subgoal_tac "inverse (z - x) * (f z - f x - f' x * (z - x)) = \ |
480 |
\ (f z - f x)/(z - x) - f' x" 2); |
|
14334 | 481 |
by (asm_full_simp_tac (simpset() addsimps [abs_mult RS sym] @ mult_ac) 2); |
13958 | 482 |
by (asm_full_simp_tac (simpset() addsimps [real_diff_def]) 2); |
483 |
by (rtac (real_mult_commute RS subst) 2); |
|
14334 | 484 |
by (asm_full_simp_tac (simpset() addsimps [left_distrib, |
13958 | 485 |
real_diff_def]) 2); |
486 |
by (dtac (CLAIM "z ~= x ==> z + -x ~= (0::real)") 2); |
|
487 |
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc, |
|
488 |
real_divide_def]) 2); |
|
14334 | 489 |
by (simp_tac (simpset() addsimps [left_distrib]) 2); |
13958 | 490 |
by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac); |
491 |
by (res_inst_tac [("u1","u"),("v1","v")] (ARITH_PROVE "u < v | v <= (u::real)" |
|
492 |
RS disjE) 1); |
|
493 |
by (dres_inst_tac [("x","v")] real_le_imp_less_or_eq 2); |
|
494 |
by Auto_tac; |
|
495 |
(*29*) |
|
496 |
by (res_inst_tac [("j","abs((f(v) - f(x)) - (f'(x) * (v - x))) + \ |
|
497 |
\ abs((f(x) - f(u)) - (f'(x) * (x - u)))")] real_le_trans 1); |
|
498 |
by (rtac (abs_triangle_ineq RSN (2,real_le_trans)) 1); |
|
14334 | 499 |
by (asm_full_simp_tac (simpset() addsimps [right_diff_distrib]) 1); |
13958 | 500 |
by (arith_tac 1); |
501 |
by (res_inst_tac [("t","e*(v - u)")] (real_sum_of_halves RS subst) 1); |
|
14334 | 502 |
by (rtac add_mono 1); |
13958 | 503 |
by (res_inst_tac [("j","(e / 2) * abs(v - x)")] real_le_trans 1); |
504 |
||
505 |
by (Asm_full_simp_tac 2 THEN arith_tac 2); |
|
506 |
by (ALLGOALS (thin_tac "ALL xa. \ |
|
507 |
\ xa ~= x & abs (xa + - x) < s --> \ |
|
508 |
\ abs ((f xa + - f x) / (xa + - x) + - f' x) * 2 < e")); |
|
509 |
by (dres_inst_tac [("x","v")] spec 1 THEN Auto_tac); |
|
510 |
by (arith_tac 1); |
|
511 |
by (dres_inst_tac [("x","u")] spec 1 THEN Auto_tac); |
|
512 |
by (arith_tac 1); |
|
513 |
by (subgoal_tac "abs (f u - f x - f' x * (u - x)) = \ |
|
514 |
\ abs (f x - f u - f' x * (x - u))" 1); |
|
515 |
by (Asm_full_simp_tac 1); |
|
14334 | 516 |
by (asm_full_simp_tac (simpset() addsimps [right_distrib, |
13958 | 517 |
real_diff_def]) 2); |
518 |
by (arith_tac 2); |
|
519 |
by(rtac real_le_trans 1); |
|
520 |
by Auto_tac; |
|
521 |
by (arith_tac 1); |
|
522 |
qed "lemma_straddle"; |
|
523 |
||
524 |
Goal "((number_of c - 1) * x <= 0) =((number_of c ::real) * x <= x)"; |
|
525 |
by (rtac ((ARITH_PROVE "(x <= y) = (x - y <= (0::real))") RS ssubst) 1); |
|
14334 | 526 |
by (simp_tac (simpset() addsimps [left_diff_distrib, |
527 |
CLAIM_SIMP "c * x - x = (c - 1) * (x::real)" [left_diff_distrib]]) 1); |
|
13958 | 528 |
qed "lemma_number_of_mult_le"; |
529 |
||
530 |
||
531 |
Goal "[|a <= b; ALL x. a <= x & x <= b --> DERIV f x :> f'(x) |] \ |
|
532 |
\ ==> Integral(a,b) f' (f(b) - f(a))"; |
|
533 |
by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac); |
|
534 |
by (auto_tac (claset(),simpset() addsimps [Integral_def])); |
|
535 |
by (rtac ccontr 1); |
|
536 |
by (subgoal_tac "ALL e. 0 < e --> \ |
|
537 |
\ (EX g. gauge (%x. a <= x & x <= b) g & \ |
|
538 |
\ (ALL D p. \ |
|
539 |
\ tpart (a, b) (D, p) & fine g (D, p) --> \ |
|
540 |
\ abs (rsum (D, p) f' - (f b - f a)) <= e))" 1); |
|
541 |
by (rotate_tac 3 1); |
|
542 |
by (dres_inst_tac [("x","e/2")] spec 1 THEN Auto_tac); |
|
543 |
by (dtac spec 1 THEN Auto_tac); |
|
544 |
by (REPEAT(dtac spec 1) THEN Auto_tac); |
|
545 |
by (dres_inst_tac [("e","ea/(b - a)")] lemma_straddle 1); |
|
14277
ad66687ece6e
more field division lemmas transferred from Real to Ring_and_Field
paulson
parents:
13958
diff
changeset
|
546 |
by (auto_tac (claset(),simpset() addsimps [thm"Ring_and_Field.zero_less_divide_iff"])); |
13958 | 547 |
by (rtac exI 1); |
548 |
by (auto_tac (claset(),simpset() addsimps [tpart_def,rsum_def])); |
|
549 |
by (subgoal_tac "sumr 0 (psize D) (%n. f(D(Suc n)) - f(D n)) = \ |
|
550 |
\ f b - f a" 1); |
|
551 |
by (cut_inst_tac [("D","%n. f(D n)"),("m","psize D")] |
|
552 |
sumr_partition_eq_diff_bounds 2); |
|
553 |
by (asm_full_simp_tac (simpset() addsimps [partition_lhs,partition_rhs]) 2); |
|
554 |
by (dtac sym 1 THEN Asm_full_simp_tac 1); |
|
555 |
by (simp_tac (simpset() addsimps [sumr_diff]) 1); |
|
556 |
by (rtac (sumr_rabs RS real_le_trans) 1); |
|
557 |
by (subgoal_tac |
|
558 |
"ea = sumr 0 (psize D) (%n. (ea / (b - a)) * (D(Suc n) - (D n)))" 1); |
|
559 |
by (asm_full_simp_tac (simpset() addsimps |
|
560 |
[ARITH_PROVE "abs(y - (x::real)) = abs(x - y)"]) 1); |
|
561 |
by (res_inst_tac [("t","ea")] ssubst 1 THEN assume_tac 1); |
|
562 |
by (rtac sumr_le2 1); |
|
563 |
by (rtac (sumr_mult RS subst) 2); |
|
564 |
by (auto_tac (claset(),simpset() addsimps [partition_rhs, |
|
565 |
partition_lhs,partition_lb,partition_ub,fine_def])); |
|
566 |
qed "FTC1"; |
|
567 |
||
568 |
||
569 |
Goal "[| Integral(a,b) f k1; k2 = k1 |] ==> Integral(a,b) f k2"; |
|
570 |
by (Asm_simp_tac 1); |
|
571 |
qed "Integral_subst"; |
|
572 |
||
573 |
Goal "[| a <= b; b <= c; Integral(a,b) f' k1; Integral(b,c) f' k2; \ |
|
574 |
\ ALL x. a <= x & x <= c --> DERIV f x :> f' x |] \ |
|
575 |
\ ==> Integral(a,c) f' (k1 + k2)"; |
|
576 |
by (rtac (FTC1 RS Integral_subst) 1); |
|
577 |
by Auto_tac; |
|
578 |
by (ftac FTC1 1 THEN Auto_tac); |
|
579 |
by (forw_inst_tac [("a","b")] FTC1 1 THEN Auto_tac); |
|
580 |
by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac); |
|
581 |
by (dres_inst_tac [("k2.0","f b - f a")] Integral_unique 1); |
|
582 |
by (dres_inst_tac [("k2.0","f c - f b")] Integral_unique 3); |
|
583 |
by Auto_tac; |
|
584 |
qed "Integral_add"; |
|
585 |
||
586 |
Goal "partition(a,b) D ==> psize D = (LEAST n. D(n) = b)"; |
|
587 |
by (auto_tac (claset() addSIs [Least_equality RS sym,partition_rhs],simpset())); |
|
588 |
by (rtac ccontr 1); |
|
589 |
by (dtac partition_ub_lt 1); |
|
590 |
by Auto_tac; |
|
591 |
qed "partition_psize_Least"; |
|
592 |
||
593 |
Goal "partition (a, c) D ==> ~ (EX n. c < D(n))"; |
|
594 |
by (Step_tac 1); |
|
595 |
by (dres_inst_tac [("r","n")] partition_ub 1); |
|
596 |
by Auto_tac; |
|
597 |
qed "lemma_partition_bounded"; |
|
598 |
||
599 |
Goal "partition (a, c) D ==> D = (%n. if D n < c then D n else c)"; |
|
600 |
by (rtac ext 1 THEN Auto_tac); |
|
601 |
by (auto_tac (claset() addSDs [lemma_partition_bounded],simpset())); |
|
602 |
by (dres_inst_tac [("x","n")] spec 1); |
|
603 |
by Auto_tac; |
|
604 |
qed "lemma_partition_eq"; |
|
605 |
||
606 |
Goal "partition (a, c) D ==> D = (%n. if D n <= c then D n else c)"; |
|
607 |
by (rtac ext 1 THEN Auto_tac); |
|
608 |
by (auto_tac (claset() addSDs [lemma_partition_bounded],simpset())); |
|
609 |
by (dres_inst_tac [("x","n")] spec 1); |
|
610 |
by Auto_tac; |
|
611 |
qed "lemma_partition_eq2"; |
|
612 |
||
613 |
Goal "[| partition(a,b) D; n < psize D |] ==> D n < D (Suc n)"; |
|
614 |
by (auto_tac (claset(),simpset() addsimps [partition])); |
|
615 |
qed "partition_lt_Suc"; |
|
616 |
||
617 |
Goal "tpart(a,c) (D,p) ==> p = (%n. if D n < c then p n else c)"; |
|
618 |
by (rtac ext 1); |
|
619 |
by (auto_tac (claset(),simpset() addsimps [tpart_def])); |
|
14334 | 620 |
by (dtac (linorder_not_less RS iffD1) 1); |
13958 | 621 |
by (dres_inst_tac [("r","Suc n")] partition_ub 1); |
622 |
by (dres_inst_tac [("x","n")] spec 1); |
|
623 |
by Auto_tac; |
|
624 |
qed "tpart_tag_eq"; |
|
625 |
||
626 |
(*----------------------------------------------------------------------------*) |
|
627 |
(* Lemmas for Additivity Theorem of gauge integral *) |
|
628 |
(*----------------------------------------------------------------------------*) |
|
629 |
||
630 |
Goal "[| a <= D n; D n < b; partition(a,b) D |] ==> n < psize D"; |
|
631 |
by (dtac (partition RS iffD1) 1 THEN Step_tac 1); |
|
632 |
by (rtac ccontr 1 THEN dtac leI 1); |
|
633 |
by Auto_tac; |
|
634 |
val lemma_additivity1 = result(); |
|
635 |
||
636 |
Goal "[| a <= D n; partition(a,D n) D |] ==> psize D <= n"; |
|
637 |
by (rtac ccontr 1 THEN dtac not_leE 1); |
|
638 |
by (ftac (partition RS iffD1) 1 THEN Step_tac 1); |
|
639 |
by (forw_inst_tac [("r","Suc n")] partition_ub 1); |
|
640 |
by (auto_tac (claset() addSDs [spec],simpset())); |
|
641 |
val lemma_additivity2 = result(); |
|
642 |
||
643 |
Goal "[| partition(a,b) D; psize D < m |] ==> D(m) = D(psize D)"; |
|
644 |
by (auto_tac (claset(),simpset() addsimps [partition])); |
|
645 |
qed "partition_eq_bound"; |
|
646 |
||
647 |
Goal "[| partition(a,b) D; psize D < m |] ==> D(r) <= D(m)"; |
|
648 |
by (ftac (partition RS iffD1) 1 THEN Auto_tac); |
|
649 |
by (etac partition_ub 1); |
|
650 |
qed "partition_ub2"; |
|
651 |
||
652 |
Goalw [tpart_def] |
|
653 |
"[| tpart(a,b) (D,p); psize D <= m |] ==> p(m) = D(m)"; |
|
654 |
by Auto_tac; |
|
655 |
by (dres_inst_tac [("x","m")] spec 1); |
|
656 |
by (auto_tac (claset(),simpset() addsimps [partition_rhs2])); |
|
657 |
qed "tag_point_eq_partition_point"; |
|
658 |
||
659 |
Goal "[| partition(a,b) D; D m < D n |] ==> m < n"; |
|
660 |
by (cut_inst_tac [("m","n"),("n","psize D")] less_linear 1); |
|
661 |
by Auto_tac; |
|
662 |
by (rtac ccontr 1 THEN dtac leI 1 THEN dtac le_imp_less_or_eq 1); |
|
663 |
by (cut_inst_tac [("m","m"),("n","psize D")] less_linear 1); |
|
664 |
by (auto_tac (claset() addDs [partition_gt],simpset())); |
|
665 |
by (dres_inst_tac [("n","m")] partition_lt_gen 1 THEN Auto_tac); |
|
666 |
by (ftac partition_eq_bound 1); |
|
667 |
by (dtac partition_gt 2); |
|
668 |
by Auto_tac; |
|
669 |
by (rtac ccontr 1 THEN dtac leI 1 THEN dtac le_imp_less_or_eq 1); |
|
670 |
by (auto_tac (claset() addDs [partition_eq_bound],simpset())); |
|
671 |
by (rtac ccontr 1 THEN dtac leI 1 THEN dtac le_imp_less_or_eq 1); |
|
672 |
by (ftac partition_eq_bound 1 THEN assume_tac 1); |
|
673 |
by (dres_inst_tac [("m","m")] partition_eq_bound 1); |
|
674 |
by Auto_tac; |
|
675 |
qed "partition_lt_cancel"; |
|
676 |
||
677 |
Goalw [psize_def] |
|
678 |
"[| a <= D n; D n < b; partition (a, b) D |] \ |
|
679 |
\ ==> psize (%x. if D x < D n then D(x) else D n) = n"; |
|
680 |
by (ftac lemma_additivity1 1); |
|
681 |
by (assume_tac 1 THEN assume_tac 1); |
|
682 |
by (rtac some_equality 1); |
|
683 |
by (auto_tac (claset() addIs [partition_lt_Suc],simpset())); |
|
684 |
by (dres_inst_tac [("n","n")] partition_lt_gen 1); |
|
685 |
by (assume_tac 1 THEN arith_tac 1 THEN arith_tac 1); |
|
686 |
by (cut_inst_tac [("m","na"),("n","psize D")] less_linear 1); |
|
687 |
by (auto_tac (claset() addDs [partition_lt_cancel],simpset())); |
|
688 |
by (rtac ccontr 1); |
|
689 |
by (dtac (ARITH_PROVE "m ~= (n::nat) ==> m < n | n < m") 1); |
|
690 |
by (etac disjE 1); |
|
691 |
by (rotate_tac 5 1); |
|
692 |
by (dres_inst_tac [("x","n")] spec 1); |
|
693 |
by Auto_tac; |
|
694 |
by (dres_inst_tac [("n","n")] partition_lt_gen 1); |
|
695 |
by Auto_tac; |
|
696 |
by (arith_tac 1); |
|
697 |
by (dres_inst_tac [("x","n")] spec 1); |
|
698 |
by (asm_full_simp_tac (simpset() addsimps [split_if]) 1); |
|
699 |
qed "lemma_additivity4_psize_eq"; |
|
700 |
||
701 |
Goal "partition (a, b) D \ |
|
702 |
\ ==> psize (%x. if D x < D n then D(x) else D n) <= psize D"; |
|
703 |
by (forw_inst_tac [("r","n")] partition_ub 1); |
|
704 |
by (dres_inst_tac [("x","D n")] real_le_imp_less_or_eq 1); |
|
705 |
by (auto_tac (claset(),simpset() addsimps |
|
706 |
[lemma_partition_eq RS sym])); |
|
707 |
by (forw_inst_tac [("r","n")] partition_lb 1); |
|
708 |
by (dtac lemma_additivity4_psize_eq 1); |
|
709 |
by (rtac ccontr 3 THEN Auto_tac); |
|
710 |
by (ftac (not_leE RSN (2,partition_eq_bound)) 1); |
|
711 |
by (auto_tac (claset(),simpset() addsimps [partition_rhs])); |
|
712 |
qed "lemma_psize_left_less_psize"; |
|
713 |
||
714 |
Goal "[| partition(a,b) D; na < psize (%x. if D x < D n then D(x) else D n) |] \ |
|
715 |
\ ==> na < psize D"; |
|
716 |
by (etac (lemma_psize_left_less_psize RSN (2,less_le_trans)) 1); |
|
717 |
by (assume_tac 1); |
|
718 |
qed "lemma_psize_left_less_psize2"; |
|
719 |
||
720 |
||
721 |
Goal "[| partition(a,b) D; D na < D n; D n < D (Suc na); \ |
|
722 |
\ n < psize D |] \ |
|
723 |
\ ==> False"; |
|
724 |
by (cut_inst_tac [("m","n"),("n","Suc na")] less_linear 1); |
|
725 |
by Auto_tac; |
|
726 |
by (dres_inst_tac [("n","n")] partition_lt_gen 2); |
|
727 |
by Auto_tac; |
|
728 |
by (cut_inst_tac [("m","psize D"),("n","na")] less_linear 1); |
|
729 |
by (auto_tac (claset(),simpset() addsimps [partition_rhs2])); |
|
730 |
by (dtac (ARITH_PROVE "n < Suc na ==> n < na | n = na") 1); |
|
731 |
by Auto_tac; |
|
732 |
by (dres_inst_tac [("n","na")] partition_lt_gen 1); |
|
733 |
by Auto_tac; |
|
734 |
val lemma_additivity3 = result(); |
|
735 |
||
736 |
Goalw [psize_def] "psize (%x. k) = 0"; |
|
737 |
by Auto_tac; |
|
738 |
qed "psize_const"; |
|
739 |
Addsimps [psize_const]; |
|
740 |
||
741 |
Goal "[| partition(a,b) D; D na < D n; D n < D (Suc na); \ |
|
742 |
\ na < psize D |] \ |
|
743 |
\ ==> False"; |
|
744 |
by (forw_inst_tac [("m","n")] partition_lt_cancel 1); |
|
745 |
by (auto_tac (claset() addIs [lemma_additivity3],simpset())); |
|
746 |
val lemma_additivity3a = result(); |
|
747 |
||
748 |
Goal "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) <= psize D - n"; |
|
749 |
by (res_inst_tac [("D2","(%x. D (x + n))")] (psize_def RS |
|
750 |
meta_eq_to_obj_eq RS ssubst) 1); |
|
751 |
by (res_inst_tac [("a","psize D - n")] someI2 1); |
|
752 |
by Auto_tac; |
|
753 |
by (dtac (ARITH_PROVE "(m::nat) < n - p ==> m + p < n") 1); |
|
754 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
755 |
by (dtac (ARITH_PROVE "(m::nat) - n <= p ==> m <= p + n") 1); |
|
756 |
by (case_tac "psize D <= n" 1); |
|
757 |
by (dtac not_leE 2); |
|
758 |
by (asm_simp_tac (simpset() addsimps |
|
759 |
[ARITH_PROVE "p <= n ==> p - n = (0::nat)"]) 1); |
|
760 |
by (blast_tac (claset() addDs [partition_rhs2]) 1); |
|
761 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
762 |
by (rtac ccontr 1 THEN dtac not_leE 1); |
|
763 |
by (dres_inst_tac [("x","psize D - n")] spec 1); |
|
764 |
by Auto_tac; |
|
765 |
by (ftac partition_rhs 1 THEN Step_tac 1); |
|
766 |
by (ftac partition_lt_cancel 1 THEN assume_tac 1); |
|
767 |
by (dtac (partition RS iffD1) 1 THEN Step_tac 1); |
|
768 |
by (subgoal_tac "~ D (psize D - n + n) < D (Suc (psize D - n + n))" 1); |
|
769 |
by (Blast_tac 1); |
|
770 |
by (rotate_tac 6 1 THEN dres_inst_tac [("x","Suc (psize D)")] spec 1); |
|
771 |
by (Asm_simp_tac 1); |
|
772 |
qed "better_lemma_psize_right_eq1"; |
|
773 |
||
774 |
Goal "partition (a, D n) D ==> psize D <= n"; |
|
775 |
by (rtac ccontr 1 THEN dtac not_leE 1); |
|
776 |
by (ftac partition_lt_Suc 1 THEN assume_tac 1); |
|
777 |
by (forw_inst_tac [("r","Suc n")] partition_ub 1); |
|
778 |
by Auto_tac; |
|
779 |
qed "psize_le_n"; |
|
780 |
||
781 |
Goal "partition(a,D n) D ==> psize (%x. D (x + n)) <= psize D - n"; |
|
782 |
by (res_inst_tac [("D2","(%x. D (x + n))")] (psize_def RS |
|
783 |
meta_eq_to_obj_eq RS ssubst) 1); |
|
784 |
by (res_inst_tac [("a","psize D - n")] someI2 1); |
|
785 |
by Auto_tac; |
|
786 |
by (dtac (ARITH_PROVE "(m::nat) < n - p ==> m + p < n") 1); |
|
787 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
788 |
by (dtac (ARITH_PROVE "(m::nat) - n <= p ==> m <= p + n") 1); |
|
789 |
by (case_tac "psize D <= n" 1); |
|
790 |
by (dtac not_leE 2); |
|
791 |
by (asm_simp_tac (simpset() addsimps |
|
792 |
[ARITH_PROVE "p <= n ==> p - n = (0::nat)"]) 1); |
|
793 |
by (blast_tac (claset() addDs [partition_rhs2]) 1); |
|
794 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
795 |
by (rtac ccontr 1 THEN dtac not_leE 1); |
|
796 |
by (ftac psize_le_n 1); |
|
797 |
by (dres_inst_tac [("x","psize D - n")] spec 1); |
|
798 |
by (asm_full_simp_tac (simpset() addsimps |
|
799 |
[ARITH_PROVE "p <= n ==> p - n = (0::nat)"]) 1); |
|
800 |
by (dtac (partition RS iffD1) 1 THEN Step_tac 1); |
|
801 |
by (rotate_tac 5 1 THEN dres_inst_tac [("x","Suc n")] spec 1); |
|
802 |
by Auto_tac; |
|
803 |
qed "better_lemma_psize_right_eq1a"; |
|
804 |
||
805 |
Goal "partition(a,b) D ==> psize (%x. D (x + n)) <= psize D - n"; |
|
806 |
by (forw_inst_tac [("r1","n")] (partition_ub RS real_le_imp_less_or_eq) 1); |
|
807 |
by (blast_tac (claset() addIs [better_lemma_psize_right_eq1a, |
|
808 |
better_lemma_psize_right_eq1]) 1); |
|
809 |
qed "better_lemma_psize_right_eq"; |
|
810 |
||
811 |
Goal "[| partition(a,b) D; D n < b |] ==> psize (%x. D (x + n)) <= psize D"; |
|
812 |
by (res_inst_tac [("D2","(%x. D (x + n))")] (psize_def RS |
|
813 |
meta_eq_to_obj_eq RS ssubst) 1); |
|
814 |
by (res_inst_tac [("a","psize D - n")] someI2 1); |
|
815 |
by Auto_tac; |
|
816 |
by (dtac (ARITH_PROVE "(m::nat) < n - p ==> m + p < n") 1); |
|
817 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
818 |
by (subgoal_tac "n <= psize D" 1); |
|
819 |
by (dtac (ARITH_PROVE "(m::nat) - n <= p ==> m <= p + n") 1); |
|
820 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
821 |
by (rtac ccontr 1 THEN dtac not_leE 1); |
|
822 |
by (dtac (less_imp_le RSN (2,partition_rhs2)) 1); |
|
823 |
by Auto_tac; |
|
824 |
by (rtac ccontr 1 THEN dtac not_leE 1); |
|
825 |
by (dres_inst_tac [("x","psize D")] spec 1); |
|
826 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
827 |
qed "lemma_psize_right_eq1"; |
|
828 |
||
829 |
(* should be combined with previous theorem; also proof has redundancy *) |
|
830 |
Goal "partition(a,D n) D ==> psize (%x. D (x + n)) <= psize D"; |
|
831 |
by (res_inst_tac [("D2","(%x. D (x + n))")] (psize_def RS |
|
832 |
meta_eq_to_obj_eq RS ssubst) 1); |
|
833 |
by (res_inst_tac [("a","psize D - n")] someI2 1); |
|
834 |
by Auto_tac; |
|
835 |
by (dtac (ARITH_PROVE "(m::nat) < n - p ==> m + p < n") 1); |
|
836 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
837 |
by (dtac (ARITH_PROVE "(m::nat) - n <= p ==> m <= p + n") 1); |
|
838 |
by (case_tac "psize D <= n" 1); |
|
839 |
by (dtac not_leE 2); |
|
840 |
by (asm_simp_tac (simpset() addsimps |
|
841 |
[ARITH_PROVE "p <= n ==> p - n = (0::nat)"]) 1); |
|
842 |
by (blast_tac (claset() addDs [partition_rhs2]) 1); |
|
843 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
844 |
by (rtac ccontr 1 THEN dtac not_leE 1); |
|
845 |
by (dres_inst_tac [("x","psize D")] spec 1); |
|
846 |
by (asm_full_simp_tac (simpset() addsimps [partition]) 1); |
|
847 |
qed "lemma_psize_right_eq1a"; |
|
848 |
||
849 |
Goal "[| partition(a,b) D |] ==> psize (%x. D (x + n)) <= psize D"; |
|
850 |
by (forw_inst_tac [("r1","n")] (partition_ub RS real_le_imp_less_or_eq) 1); |
|
851 |
by (blast_tac (claset() addIs [lemma_psize_right_eq1a,lemma_psize_right_eq1]) 1); |
|
852 |
qed "lemma_psize_right_eq"; |
|
853 |
||
854 |
Goal "[| a <= D n; tpart (a, b) (D, p) |] \ |
|
855 |
\ ==> tpart(a, D n) (%x. if D x < D n then D(x) else D n, \ |
|
856 |
\ %x. if D x < D n then p(x) else D n)"; |
|
857 |
by (forw_inst_tac [("r","n")] (tpart_partition RS partition_ub) 1); |
|
858 |
by (dres_inst_tac [("x","D n")] real_le_imp_less_or_eq 1); |
|
859 |
by (auto_tac (claset(),simpset() addsimps |
|
860 |
[tpart_partition RS lemma_partition_eq RS sym, |
|
861 |
tpart_tag_eq RS sym])); |
|
862 |
by (ftac (tpart_partition RSN (3,lemma_additivity1)) 1); |
|
863 |
by (auto_tac (claset(),simpset() addsimps [tpart_def])); |
|
14334 | 864 |
by (dtac ((linorder_not_less RS iffD1) RS real_le_imp_less_or_eq) 2); |
13958 | 865 |
by (Auto_tac); |
866 |
by (blast_tac (claset() addDs [lemma_additivity3]) 2); |
|
14334 | 867 |
by (dtac (linorder_not_less RS iffD1) 2 THEN dres_inst_tac [("x","na")] spec 2); |
13958 | 868 |
by (arith_tac 2); |
869 |
by (ftac lemma_additivity4_psize_eq 1); |
|
870 |
by (REPEAT(assume_tac 1)); |
|
871 |
by (rtac (partition RS iffD2) 1); |
|
872 |
by (ftac (partition RS iffD1) 1); |
|
873 |
by (auto_tac (claset() addIs [partition_lt_gen],simpset() addsimps [])); |
|
874 |
by (dres_inst_tac [("n","n")] partition_lt_gen 1); |
|
875 |
by (assume_tac 1 THEN arith_tac 1); |
|
876 |
by (Blast_tac 1); |
|
877 |
by (dtac partition_lt_cancel 1 THEN Auto_tac); |
|
878 |
qed "tpart_left1"; |
|
879 |
||
880 |
Goal "[| a <= D n; tpart (a, b) (D, p); gauge (%x. a <= x & x <= D n) g;\ |
|
881 |
\ fine (%x. if x < D n then min (g x) ((D n - x)/ 2) \ |
|
882 |
\ else if x = D n then min (g (D n)) (ga (D n)) \ |
|
883 |
\ else min (ga x) ((x - D n)/ 2)) (D, p) |] \ |
|
884 |
\ ==> fine g \ |
|
885 |
\ (%x. if D x < D n then D(x) else D n, \ |
|
886 |
\ %x. if D x < D n then p(x) else D n)"; |
|
887 |
by (auto_tac (claset(),simpset() addsimps [fine_def,tpart_def,gauge_def])); |
|
888 |
by (ALLGOALS (ftac lemma_psize_left_less_psize2)); |
|
889 |
by (TRYALL(assume_tac)); |
|
890 |
by (ALLGOALS (dres_inst_tac [("x","na")] spec) THEN Auto_tac); |
|
891 |
by (ALLGOALS(dres_inst_tac [("x","na")] spec)); |
|
892 |
by Auto_tac; |
|
893 |
by (asm_full_simp_tac (simpset() addsimps [split_if]) 1); |
|
14334 | 894 |
by (dtac ((linorder_not_less RS iffD1) RS real_le_imp_less_or_eq) 1); |
13958 | 895 |
by (Step_tac 1); |
896 |
by (blast_tac (claset() addDs [lemma_additivity3a]) 1); |
|
897 |
by (dtac sym 1 THEN Auto_tac); |
|
898 |
qed "fine_left1"; |
|
899 |
||
900 |
Goal "[| a <= D n; tpart (a, b) (D, p) |] \ |
|
901 |
\ ==> tpart(D n, b) (%x. D(x + n),%x. p(x + n))"; |
|
902 |
by (asm_full_simp_tac (simpset() addsimps [tpart_def,partition_def]) 1); |
|
903 |
by (Step_tac 1); |
|
904 |
by (res_inst_tac [("x","N - n")] exI 1 THEN Auto_tac); |
|
905 |
by (rotate_tac 2 1); |
|
906 |
by (dres_inst_tac [("x","na + n")] spec 1); |
|
907 |
by (rotate_tac 3 2); |
|
908 |
by (dres_inst_tac [("x","na + n")] spec 2); |
|
909 |
by (ALLGOALS(arith_tac)); |
|
910 |
qed "tpart_right1"; |
|
911 |
||
912 |
Goal "[| a <= D n; tpart (a, b) (D, p); gauge (%x. D n <= x & x <= b) ga;\ |
|
913 |
\ fine (%x. if x < D n then min (g x) ((D n - x)/ 2) \ |
|
914 |
\ else if x = D n then min (g (D n)) (ga (D n)) \ |
|
915 |
\ else min (ga x) ((x - D n)/ 2)) (D, p) |] \ |
|
916 |
\ ==> fine ga (%x. D(x + n),%x. p(x + n))"; |
|
917 |
by (auto_tac (claset(),simpset() addsimps [fine_def,gauge_def])); |
|
918 |
by (dres_inst_tac [("x","na + n")] spec 1); |
|
919 |
by (forw_inst_tac [("n","n")] (tpart_partition RS better_lemma_psize_right_eq) 1); |
|
920 |
by Auto_tac; |
|
921 |
by (arith_tac 1); |
|
922 |
by (asm_full_simp_tac (simpset() addsimps [tpart_def]) 1 THEN Step_tac 1); |
|
923 |
by (subgoal_tac "D n <= p (na + n)" 1); |
|
924 |
by (dres_inst_tac [("y","p (na + n)")] real_le_imp_less_or_eq 1); |
|
925 |
by (Step_tac 1); |
|
926 |
by (asm_full_simp_tac (simpset() addsimps [split_if]) 1); |
|
927 |
by (Asm_full_simp_tac 1); |
|
928 |
by (dtac less_le_trans 1 THEN assume_tac 1); |
|
929 |
by (rotate_tac 5 1); |
|
930 |
by (dres_inst_tac [("x","na + n")] spec 1 THEN Step_tac 1); |
|
931 |
by (rtac real_le_trans 1 THEN assume_tac 2); |
|
932 |
by (case_tac "na = 0" 1 THEN Auto_tac); |
|
933 |
by (rtac (partition_lt_gen RS order_less_imp_le) 1); |
|
934 |
by Auto_tac; |
|
935 |
by (arith_tac 1); |
|
936 |
qed "fine_right1"; |
|
937 |
||
938 |
Goalw [rsum_def] |
|
939 |
"rsum (D, p) (%x. f x + g x) = rsum (D, p) f + rsum(D, p) g"; |
|
14334 | 940 |
by (auto_tac (claset(),simpset() addsimps [sumr_add,left_distrib])); |
13958 | 941 |
qed "rsum_add"; |
942 |
||
943 |
(* Bartle/Sherbert: Theorem 10.1.5 p. 278 *) |
|
944 |
Goalw [Integral_def] |
|
945 |
"[| a <= b; Integral(a,b) f k1; Integral(a,b) g k2 |] \ |
|
946 |
\ ==> Integral(a,b) (%x. f x + g x) (k1 + k2)"; |
|
947 |
by Auto_tac; |
|
948 |
by (REPEAT(dres_inst_tac [("x","e/2")] spec 1)); |
|
949 |
by Auto_tac; |
|
950 |
by (dtac gauge_min 1 THEN assume_tac 1); |
|
951 |
by (res_inst_tac [("x","(%x. if ga x < gaa x then ga x else gaa x)")] exI 1); |
|
952 |
by Auto_tac; |
|
953 |
by (dtac fine_min 1); |
|
954 |
by (REPEAT(dtac spec 1) THEN Auto_tac); |
|
14334 | 955 |
by (dres_inst_tac [("a","abs (rsum (D, p) f - k1)* 2"), |
956 |
("c","abs (rsum (D, p) g - k2) * 2")] |
|
957 |
add_strict_mono 1 THEN assume_tac 1); |
|
958 |
by (auto_tac (claset(),HOL_ss addsimps [rsum_add,left_distrib RS sym, |
|
13958 | 959 |
real_mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"])); |
960 |
by (arith_tac 1); |
|
961 |
qed "Integral_add_fun"; |
|
962 |
||
963 |
Goal "[| partition(a,b) D; r < psize D |] ==> 0 < D (Suc r) - D r"; |
|
964 |
by (auto_tac (claset(),simpset() addsimps [partition])); |
|
965 |
qed "partition_lt_gen2"; |
|
966 |
||
967 |
Goalw [tpart_def] |
|
968 |
"[| ALL x. a <= x & x <= b --> f x <= g x; \ |
|
969 |
\ tpart(a,b) (D,p) \ |
|
970 |
\ |] ==> ALL n. n <= psize D --> f (p n) <= g (p n)"; |
|
971 |
by (Auto_tac THEN ftac (partition RS iffD1) 1); |
|
972 |
by Auto_tac; |
|
973 |
by (dres_inst_tac [("x","p n")] spec 1); |
|
974 |
by Auto_tac; |
|
975 |
by (case_tac "n = 0" 1); |
|
976 |
by (rtac (partition_lt_gen RS order_less_le_trans RS order_less_imp_le) 2); |
|
977 |
by Auto_tac; |
|
978 |
by (dtac le_imp_less_or_eq 1 THEN Auto_tac); |
|
979 |
by (dres_inst_tac [("x","psize D")] spec 2 THEN Auto_tac); |
|
980 |
by (dres_inst_tac [("r","Suc n")] partition_ub 1); |
|
981 |
by (dres_inst_tac [("x","n")] spec 1); |
|
982 |
by Auto_tac; |
|
983 |
qed "lemma_Integral_le"; |
|
984 |
||
985 |
Goalw [rsum_def] |
|
986 |
"[| ALL x. a <= x & x <= b --> f x <= g x; \ |
|
987 |
\ tpart(a,b) (D,p) \ |
|
988 |
\ |] ==> rsum(D,p) f <= rsum(D,p) g"; |
|
989 |
by (auto_tac (claset() addSIs [sumr_le2] addDs |
|
990 |
[tpart_partition RS partition_lt_gen2] addSDs |
|
991 |
[lemma_Integral_le],simpset())); |
|
992 |
qed "lemma_Integral_rsum_le"; |
|
993 |
||
994 |
Goalw [Integral_def] |
|
995 |
"[| a <= b; \ |
|
996 |
\ ALL x. a <= x & x <= b --> f(x) <= g(x); \ |
|
997 |
\ Integral(a,b) f k1; Integral(a,b) g k2 \ |
|
998 |
\ |] ==> k1 <= k2"; |
|
999 |
by Auto_tac; |
|
1000 |
by (rotate_tac 2 1); |
|
1001 |
by (dres_inst_tac [("x","abs(k1 - k2)/2")] spec 1); |
|
1002 |
by (dres_inst_tac [("x","abs(k1 - k2)/2")] spec 1); |
|
1003 |
by Auto_tac; |
|
1004 |
by (dtac gauge_min 1 THEN assume_tac 1); |
|
1005 |
by (dres_inst_tac [("g","%x. if ga x < gaa x then ga x else gaa x")] |
|
1006 |
partition_exists 1 THEN assume_tac 1); |
|
1007 |
by Auto_tac; |
|
1008 |
by (dtac fine_min 1); |
|
1009 |
by (dres_inst_tac [("x","D")] spec 1 THEN dres_inst_tac [("x","D")] spec 1); |
|
1010 |
by (dres_inst_tac [("x","p")] spec 1 THEN dres_inst_tac [("x","p")] spec 1); |
|
1011 |
by Auto_tac; |
|
1012 |
by (ftac lemma_Integral_rsum_le 1 THEN assume_tac 1); |
|
1013 |
by (subgoal_tac |
|
1014 |
"abs((rsum(D,p) f - k1) - (rsum(D,p) g - k2)) < abs(k1 - k2)" 1); |
|
1015 |
by (arith_tac 1); |
|
14334 | 1016 |
by (dtac add_strict_mono 1 THEN assume_tac 1); |
1017 |
by (auto_tac (claset(),HOL_ss addsimps [left_distrib RS sym, |
|
13958 | 1018 |
real_mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"])); |
1019 |
by (arith_tac 1); |
|
1020 |
qed "Integral_le"; |
|
1021 |
||
1022 |
Goalw [Integral_def] |
|
1023 |
"(EX k. Integral(a,b) f k) ==> \ |
|
1024 |
\ (ALL e. 0 < e --> \ |
|
1025 |
\ (EX g. gauge (%x. a <= x & x <= b) g & \ |
|
1026 |
\ (ALL D1 D2 p1 p2. \ |
|
1027 |
\ tpart(a,b) (D1, p1) & fine g (D1,p1) & \ |
|
1028 |
\ tpart(a,b) (D2, p2) & fine g (D2,p2) --> \ |
|
1029 |
\ abs(rsum(D1,p1) f - rsum(D2,p2) f) < e)))"; |
|
1030 |
by Auto_tac; |
|
1031 |
by (dres_inst_tac [("x","e/2")] spec 1); |
|
1032 |
by Auto_tac; |
|
1033 |
by (rtac exI 1 THEN Auto_tac); |
|
1034 |
by (forw_inst_tac [("x","D1")] spec 1); |
|
1035 |
by (forw_inst_tac [("x","D2")] spec 1); |
|
1036 |
by (REPEAT(dtac spec 1) THEN Auto_tac); |
|
1037 |
by (thin_tac "0 < e" 1); |
|
14334 | 1038 |
by (dtac add_strict_mono 1 THEN assume_tac 1); |
1039 |
by (auto_tac (claset(),HOL_ss addsimps [left_distrib RS sym, |
|
13958 | 1040 |
real_mult_2_right RS sym,real_mult_less_iff1,CLAIM "(0::real) < 2"])); |
1041 |
by (arith_tac 1); |
|
1042 |
qed "Integral_imp_Cauchy"; |
|
1043 |
||
1044 |
Goalw [Cauchy_def] |
|
1045 |
"Cauchy X = \ |
|
1046 |
\ (ALL j. (EX M. ALL m n. M <= m & M <= n --> \ |
|
1047 |
\ abs (X m + - X n) < inverse(real (Suc j))))"; |
|
1048 |
by Auto_tac; |
|
1049 |
by (dtac reals_Archimedean 1); |
|
1050 |
by (Step_tac 1); |
|
1051 |
by (dres_inst_tac [("x","n")] spec 1); |
|
1052 |
by Auto_tac; |
|
1053 |
by (res_inst_tac [("x","M")] exI 1); |
|
1054 |
by Auto_tac; |
|
1055 |
by (dres_inst_tac [("x","m")] spec 1); |
|
1056 |
by (dres_inst_tac [("x","na")] spec 1); |
|
1057 |
by Auto_tac; |
|
1058 |
qed "Cauchy_iff2"; |
|
1059 |
||
1060 |
Goal "[| a <= b; ALL n. gauge (%x. a <= x & x <= b) (fa n) |] \ |
|
1061 |
\ ==> ALL n. EX D p. tpart (a, b) (D, p) & fine (fa n) (D, p)"; |
|
1062 |
by (Step_tac 1); |
|
1063 |
by (rtac partition_exists 1 THEN Auto_tac); |
|
1064 |
qed "partition_exists2"; |
|
1065 |
||
1066 |
Goal "[| a <= b; ALL c. a <= c & c <= b --> f' c <= g' c; \ |
|
1067 |
\ ALL x. DERIV f x :> f' x; ALL x. DERIV g x :> g' x \ |
|
1068 |
\ |] ==> f b - f a <= g b - g a"; |
|
1069 |
by (rtac Integral_le 1); |
|
1070 |
by (assume_tac 1); |
|
1071 |
by (rtac FTC1 2); |
|
1072 |
by (rtac FTC1 4); |
|
1073 |
by Auto_tac; |
|
1074 |
qed "monotonic_anti_derivative"; |