| author | wenzelm | 
| Fri, 14 Dec 2001 11:53:31 +0100 | |
| changeset 12499 | 1b56e1732a61 | 
| parent 12486 | 0ed8bdd883e0 | 
| child 13601 | fd3e3d6b37b2 | 
| permissions | -rw-r--r-- | 
| 10751 | 1  | 
(* Title : Lim.ML  | 
2  | 
Author : Jacques D. Fleuriot  | 
|
3  | 
Copyright : 1998 University of Cambridge  | 
|
4  | 
Description : Theory of limits, continuity and  | 
|
5  | 
differentiation of real=>real functions  | 
|
6  | 
*)  | 
|
7  | 
||
8  | 
fun ARITH_PROVE str = prove_goal thy str  | 
|
9  | 
(fn prems => [cut_facts_tac prems 1,arith_tac 1]);  | 
|
10  | 
||
11  | 
||
12  | 
(*---------------------------------------------------------------  | 
|
13  | 
Theory of limits, continuity and differentiation of  | 
|
14  | 
real=>real functions  | 
|
15  | 
----------------------------------------------------------------*)  | 
|
16  | 
||
17  | 
Goalw [LIM_def] "(%x. k) -- x --> k";  | 
|
18  | 
by Auto_tac;  | 
|
19  | 
qed "LIM_const";  | 
|
20  | 
Addsimps [LIM_const];  | 
|
21  | 
||
22  | 
(***-----------------------------------------------------------***)  | 
|
23  | 
(*** Some Purely Standard Proofs - Can be used for comparison ***)  | 
|
24  | 
(***-----------------------------------------------------------***)  | 
|
25  | 
||
26  | 
(*---------------  | 
|
27  | 
LIM_add  | 
|
28  | 
---------------*)  | 
|
29  | 
Goalw [LIM_def]  | 
|
| 
10778
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
30  | 
"[| f -- x --> l; g -- x --> m |] ==> (%x. f(x) + g(x)) -- x --> (l + m)";  | 
| 
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
31  | 
by (Clarify_tac 1);  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
32  | 
by (REPEAT(dres_inst_tac [("x","r/2")] spec 1));
 | 
| 10751 | 33  | 
by (Asm_full_simp_tac 1);  | 
34  | 
by (Clarify_tac 1);  | 
|
35  | 
by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
 | 
|
36  | 
real_linear_less2 1);  | 
|
37  | 
by (res_inst_tac [("x","s")] exI 1);
 | 
|
38  | 
by (res_inst_tac [("x","sa")] exI 2);
 | 
|
39  | 
by (res_inst_tac [("x","sa")] exI 3);
 | 
|
| 11176 | 40  | 
by Safe_tac;  | 
| 10751 | 41  | 
by (REPEAT(dres_inst_tac [("x","xa")] spec 1) 
 | 
42  | 
THEN step_tac (claset() addSEs [order_less_trans]) 1);  | 
|
43  | 
by (REPEAT(dres_inst_tac [("x","xa")] spec 2) 
 | 
|
44  | 
THEN step_tac (claset() addSEs [order_less_trans]) 2);  | 
|
45  | 
by (REPEAT(dres_inst_tac [("x","xa")] spec 3) 
 | 
|
46  | 
THEN step_tac (claset() addSEs [order_less_trans]) 3);  | 
|
47  | 
by (ALLGOALS(rtac (abs_sum_triangle_ineq RS order_le_less_trans)));  | 
|
48  | 
by (ALLGOALS(rtac (real_sum_of_halves RS subst)));  | 
|
49  | 
by (auto_tac (claset() addIs [real_add_less_mono],simpset()));  | 
|
50  | 
qed "LIM_add";  | 
|
51  | 
||
52  | 
Goalw [LIM_def] "f -- a --> L ==> (%x. -f(x)) -- a --> -L";  | 
|
53  | 
by (full_simp_tac (simpset() addsimps [real_minus_add_distrib RS sym]  | 
|
54  | 
delsimps [real_minus_add_distrib, real_minus_minus]) 1);  | 
|
55  | 
qed "LIM_minus";  | 
|
56  | 
||
57  | 
(*----------------------------------------------  | 
|
58  | 
LIM_add_minus  | 
|
59  | 
----------------------------------------------*)  | 
|
60  | 
Goal "[| f -- x --> l; g -- x --> m |] \  | 
|
61  | 
\ ==> (%x. f(x) + -g(x)) -- x --> (l + -m)";  | 
|
62  | 
by (blast_tac (claset() addDs [LIM_add,LIM_minus]) 1);  | 
|
63  | 
qed "LIM_add_minus";  | 
|
64  | 
||
65  | 
(*----------------------------------------------  | 
|
66  | 
LIM_zero  | 
|
67  | 
----------------------------------------------*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
68  | 
Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> 0";  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
69  | 
by (res_inst_tac [("z1","l")] ((real_add_minus RS subst)) 1);
 | 
| 10751 | 70  | 
by (rtac LIM_add_minus 1 THEN Auto_tac);  | 
71  | 
qed "LIM_zero";  | 
|
72  | 
||
73  | 
(*--------------------------  | 
|
74  | 
Limit not zero  | 
|
75  | 
--------------------------*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
76  | 
Goalw [LIM_def] "k \\<noteq> 0 ==> ~ ((%x. k) -- x --> 0)";  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
77  | 
by (res_inst_tac [("R1.0","k"),("R2.0","0")] real_linear_less2 1);
 | 
| 10751 | 78  | 
by (auto_tac (claset(), simpset() addsimps [real_abs_def]));  | 
79  | 
by (res_inst_tac [("x","-k")] exI 1);
 | 
|
80  | 
by (res_inst_tac [("x","k")] exI 2);
 | 
|
81  | 
by Auto_tac;  | 
|
82  | 
by (ALLGOALS(dres_inst_tac [("y","s")] real_dense));
 | 
|
83  | 
by Safe_tac;  | 
|
84  | 
by (ALLGOALS(res_inst_tac [("x","r + x")] exI));
 | 
|
85  | 
by Auto_tac;  | 
|
86  | 
qed "LIM_not_zero";  | 
|
87  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
88  | 
(* [| k \\<noteq> 0; (%x. k) -- x --> 0 |] ==> R *)  | 
| 10751 | 89  | 
bind_thm("LIM_not_zeroE", LIM_not_zero RS notE);
 | 
90  | 
||
91  | 
Goal "(%x. k) -- x --> L ==> k = L";  | 
|
92  | 
by (rtac ccontr 1);  | 
|
93  | 
by (dtac LIM_zero 1);  | 
|
94  | 
by (rtac LIM_not_zeroE 1 THEN assume_tac 2);  | 
|
95  | 
by (arith_tac 1);  | 
|
96  | 
qed "LIM_const_eq";  | 
|
97  | 
||
98  | 
(*------------------------  | 
|
99  | 
Limit is Unique  | 
|
100  | 
------------------------*)  | 
|
101  | 
Goal "[| f -- x --> L; f -- x --> M |] ==> L = M";  | 
|
102  | 
by (dtac LIM_minus 1);  | 
|
103  | 
by (dtac LIM_add 1 THEN assume_tac 1);  | 
|
104  | 
by (auto_tac (claset() addSDs [LIM_const_eq RS sym], simpset()));  | 
|
105  | 
qed "LIM_unique";  | 
|
106  | 
||
107  | 
(*-------------  | 
|
108  | 
LIM_mult_zero  | 
|
109  | 
-------------*)  | 
|
| 11383 | 110  | 
Goalw [LIM_def]  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
111  | 
"[| f -- x --> 0; g -- x --> 0 |] ==> (%x. f(x)*g(x)) -- x --> 0";  | 
| 11176 | 112  | 
by Safe_tac;  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
113  | 
by (dres_inst_tac [("x","1")] spec 1);
 | 
| 10751 | 114  | 
by (dres_inst_tac [("x","r")] spec 1);
 | 
115  | 
by (cut_facts_tac [real_zero_less_one] 1);  | 
|
116  | 
by (asm_full_simp_tac (simpset() addsimps  | 
|
117  | 
[abs_mult]) 1);  | 
|
118  | 
by (Clarify_tac 1);  | 
|
119  | 
by (res_inst_tac [("R1.0","s"),("R2.0","sa")] 
 | 
|
120  | 
real_linear_less2 1);  | 
|
121  | 
by (res_inst_tac [("x","s")] exI 1);
 | 
|
122  | 
by (res_inst_tac [("x","sa")] exI 2);
 | 
|
123  | 
by (res_inst_tac [("x","sa")] exI 3);
 | 
|
| 11176 | 124  | 
by Safe_tac;  | 
| 10751 | 125  | 
by (REPEAT(dres_inst_tac [("x","xa")] spec 1) 
 | 
126  | 
THEN step_tac (claset() addSEs [order_less_trans]) 1);  | 
|
127  | 
by (REPEAT(dres_inst_tac [("x","xa")] spec 2) 
 | 
|
128  | 
THEN step_tac (claset() addSEs [order_less_trans]) 2);  | 
|
129  | 
by (REPEAT(dres_inst_tac [("x","xa")] spec 3) 
 | 
|
130  | 
THEN step_tac (claset() addSEs [order_less_trans]) 3);  | 
|
131  | 
by (ALLGOALS(res_inst_tac [("t","r")] (real_mult_1 RS subst)));
 | 
|
132  | 
by (ALLGOALS(rtac abs_mult_less2));  | 
|
133  | 
by Auto_tac;  | 
|
134  | 
qed "LIM_mult_zero";  | 
|
135  | 
||
136  | 
Goalw [LIM_def] "(%x. x) -- a --> a";  | 
|
137  | 
by Auto_tac;  | 
|
138  | 
qed "LIM_self";  | 
|
139  | 
||
140  | 
(*--------------------------------------------------------------  | 
|
141  | 
Limits are equal for functions equal except at limit point  | 
|
142  | 
--------------------------------------------------------------*)  | 
|
143  | 
Goalw [LIM_def]  | 
|
| 11383 | 144  | 
"[| \\<forall>x. x \\<noteq> a --> (f x = g x) |] \  | 
145  | 
\ ==> (f -- a --> l) = (g -- a --> l)";  | 
|
| 10751 | 146  | 
by (auto_tac (claset(), simpset() addsimps [real_add_minus_iff]));  | 
147  | 
qed "LIM_equal";  | 
|
148  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
149  | 
Goal "[| (%x. f(x) + -g(x)) -- a --> 0; g -- a --> l |] \  | 
| 11383 | 150  | 
\ ==> f -- a --> l";  | 
| 10751 | 151  | 
by (dtac LIM_add 1 THEN assume_tac 1);  | 
152  | 
by (auto_tac (claset(), simpset() addsimps [real_add_assoc]));  | 
|
153  | 
qed "LIM_trans";  | 
|
154  | 
||
155  | 
(***-------------------------------------------------------------***)  | 
|
156  | 
(*** End of Purely Standard Proofs ***)  | 
|
157  | 
(***-------------------------------------------------------------***)  | 
|
158  | 
(*--------------------------------------------------------------  | 
|
159  | 
Standard and NS definitions of Limit  | 
|
160  | 
--------------------------------------------------------------*)  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
161  | 
Goalw [LIM_def,NSLIM_def,approx_def]  | 
| 10751 | 162  | 
"f -- x --> L ==> f -- x --NS> L";  | 
163  | 
by (asm_full_simp_tac  | 
|
164  | 
(simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1);  | 
|
| 11176 | 165  | 
by Safe_tac;  | 
| 10751 | 166  | 
by (res_inst_tac [("z","xa")] eq_Abs_hypreal 1);
 | 
167  | 
by (auto_tac (claset(),  | 
|
168  | 
simpset() addsimps [real_add_minus_iff, starfun, hypreal_minus,  | 
|
169  | 
hypreal_of_real_def, hypreal_add]));  | 
|
170  | 
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1);  | 
|
171  | 
by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1);
 | 
|
172  | 
by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1);
 | 
|
| 11383 | 173  | 
by (subgoal_tac "\\<forall>n::nat. (xa n) \\<noteq> x & \  | 
| 10751 | 174  | 
\ abs ((xa n) + - x) < s --> abs (f (xa n) + - L) < u" 1);  | 
175  | 
by (Blast_tac 2);  | 
|
176  | 
by (dtac FreeUltrafilterNat_all 1);  | 
|
177  | 
by (Ultra_tac 1);  | 
|
178  | 
qed "LIM_NSLIM";  | 
|
179  | 
||
180  | 
(*---------------------------------------------------------------------  | 
|
181  | 
Limit: NS definition ==> standard definition  | 
|
182  | 
---------------------------------------------------------------------*)  | 
|
183  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
184  | 
Goal "\\<forall>s. 0 < s --> (\\<exists>xa. xa \\<noteq> x & \  | 
| 11383 | 185  | 
\ abs (xa + - x) < s & r \\<le> abs (f xa + -L)) \  | 
186  | 
\ ==> \\<forall>n::nat. \\<exists>xa. xa \\<noteq> x & \  | 
|
187  | 
\ abs(xa + -x) < inverse(real(Suc n)) & r \\<le> abs(f xa + -L)";  | 
|
| 
10778
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
188  | 
by (Clarify_tac 1);  | 
| 10751 | 189  | 
by (cut_inst_tac [("n1","n")]
 | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
190  | 
(real_of_nat_Suc_gt_zero RS real_inverse_gt_0) 1);  | 
| 10751 | 191  | 
by Auto_tac;  | 
192  | 
val lemma_LIM = result();  | 
|
193  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
194  | 
Goal "\\<forall>s. 0 < s --> (\\<exists>xa. xa \\<noteq> x & \  | 
| 11383 | 195  | 
\ abs (xa + - x) < s & r \\<le> abs (f xa + -L)) \  | 
196  | 
\ ==> \\<exists>X. \\<forall>n::nat. X n \\<noteq> x & \  | 
|
197  | 
\ abs(X n + -x) < inverse(real(Suc n)) & r \\<le> abs(f (X n) + -L)";  | 
|
| 10751 | 198  | 
by (dtac lemma_LIM 1);  | 
199  | 
by (dtac choice 1);  | 
|
200  | 
by (Blast_tac 1);  | 
|
201  | 
val lemma_skolemize_LIM2 = result();  | 
|
202  | 
||
| 11383 | 203  | 
Goal "\\<forall>n. X n \\<noteq> x & \  | 
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
204  | 
\ abs (X n + - x) < inverse (real(Suc n)) & \  | 
| 11383 | 205  | 
\ r \\<le> abs (f (X n) + - L) ==> \  | 
206  | 
\ \\<forall>n. abs (X n + - x) < inverse (real(Suc n))";  | 
|
| 10751 | 207  | 
by (Auto_tac );  | 
208  | 
val lemma_simp = result();  | 
|
209  | 
||
210  | 
(*-------------------  | 
|
211  | 
NSLIM => LIM  | 
|
212  | 
-------------------*)  | 
|
213  | 
||
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
214  | 
Goalw [LIM_def,NSLIM_def,approx_def]  | 
| 
10778
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
215  | 
"f -- x --NS> L ==> f -- x --> L";  | 
| 10751 | 216  | 
by (asm_full_simp_tac  | 
217  | 
(simpset() addsimps [Infinitesimal_FreeUltrafilterNat_iff]) 1);  | 
|
218  | 
by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]);  | 
|
219  | 
by (fold_tac [real_le_def]);  | 
|
220  | 
by (dtac lemma_skolemize_LIM2 1);  | 
|
| 11176 | 221  | 
by Safe_tac;  | 
| 10834 | 222  | 
by (dres_inst_tac [("x","Abs_hypreal(hyprel``{X})")] spec 1);
 | 
| 10751 | 223  | 
by (asm_full_simp_tac  | 
224  | 
(simpset() addsimps [starfun, hypreal_minus,  | 
|
225  | 
hypreal_of_real_def,hypreal_add]) 1);  | 
|
| 11176 | 226  | 
by Safe_tac;  | 
| 10751 | 227  | 
by (dtac (lemma_simp RS real_seq_to_hypreal_Infinitesimal) 1);  | 
228  | 
by (asm_full_simp_tac  | 
|
229  | 
(simpset() addsimps  | 
|
230  | 
[Infinitesimal_FreeUltrafilterNat_iff,hypreal_of_real_def,  | 
|
231  | 
hypreal_minus, hypreal_add]) 1);  | 
|
232  | 
by (Blast_tac 1);  | 
|
233  | 
by (rotate_tac 2 1);  | 
|
234  | 
by (dres_inst_tac [("x","r")] spec 1);
 | 
|
235  | 
by (Clarify_tac 1);  | 
|
236  | 
by (dtac FreeUltrafilterNat_all 1);  | 
|
237  | 
by (Ultra_tac 1);  | 
|
238  | 
qed "NSLIM_LIM";  | 
|
239  | 
||
240  | 
||
241  | 
(**** Key result ****)  | 
|
242  | 
Goal "(f -- x --> L) = (f -- x --NS> L)";  | 
|
243  | 
by (blast_tac (claset() addIs [LIM_NSLIM,NSLIM_LIM]) 1);  | 
|
244  | 
qed "LIM_NSLIM_iff";  | 
|
245  | 
||
246  | 
(*-------------------------------------------------------------------*)  | 
|
247  | 
(* Proving properties of limits using nonstandard definition and *)  | 
|
248  | 
(* hence, the properties hold for standard limits as well *)  | 
|
249  | 
(*-------------------------------------------------------------------*)  | 
|
250  | 
(*------------------------------------------------  | 
|
251  | 
NSLIM_mult and hence (trivially) LIM_mult  | 
|
252  | 
------------------------------------------------*)  | 
|
253  | 
||
254  | 
Goalw [NSLIM_def]  | 
|
255  | 
"[| f -- x --NS> l; g -- x --NS> m |] \  | 
|
256  | 
\ ==> (%x. f(x) * g(x)) -- x --NS> (l * m)";  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
257  | 
by (auto_tac (claset() addSIs [approx_mult_HFinite], simpset()));  | 
| 10751 | 258  | 
qed "NSLIM_mult";  | 
259  | 
||
260  | 
Goal "[| f -- x --> l; g -- x --> m |] \  | 
|
261  | 
\ ==> (%x. f(x) * g(x)) -- x --> (l * m)";  | 
|
262  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_mult]) 1);  | 
|
263  | 
qed "LIM_mult2";  | 
|
264  | 
||
265  | 
(*----------------------------------------------  | 
|
266  | 
NSLIM_add and hence (trivially) LIM_add  | 
|
267  | 
Note the much shorter proof  | 
|
268  | 
----------------------------------------------*)  | 
|
269  | 
Goalw [NSLIM_def]  | 
|
270  | 
"[| f -- x --NS> l; g -- x --NS> m |] \  | 
|
271  | 
\ ==> (%x. f(x) + g(x)) -- x --NS> (l + m)";  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
272  | 
by (auto_tac (claset() addSIs [approx_add], simpset()));  | 
| 10751 | 273  | 
qed "NSLIM_add";  | 
274  | 
||
275  | 
Goal "[| f -- x --> l; g -- x --> m |] \  | 
|
276  | 
\ ==> (%x. f(x) + g(x)) -- x --> (l + m)";  | 
|
277  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_add]) 1);  | 
|
278  | 
qed "LIM_add2";  | 
|
279  | 
||
280  | 
(*----------------------------------------------  | 
|
281  | 
NSLIM_const  | 
|
282  | 
----------------------------------------------*)  | 
|
283  | 
Goalw [NSLIM_def] "(%x. k) -- x --NS> k";  | 
|
284  | 
by Auto_tac;  | 
|
285  | 
qed "NSLIM_const";  | 
|
286  | 
||
287  | 
Addsimps [NSLIM_const];  | 
|
288  | 
||
289  | 
Goal "(%x. k) -- x --> k";  | 
|
290  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);  | 
|
291  | 
qed "LIM_const2";  | 
|
292  | 
||
293  | 
(*----------------------------------------------  | 
|
294  | 
NSLIM_minus  | 
|
295  | 
----------------------------------------------*)  | 
|
296  | 
Goalw [NSLIM_def]  | 
|
297  | 
"f -- a --NS> L ==> (%x. -f(x)) -- a --NS> -L";  | 
|
298  | 
by Auto_tac;  | 
|
299  | 
qed "NSLIM_minus";  | 
|
300  | 
||
301  | 
Goal "f -- a --> L ==> (%x. -f(x)) -- a --> -L";  | 
|
302  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_minus]) 1);  | 
|
303  | 
qed "LIM_minus2";  | 
|
304  | 
||
305  | 
(*----------------------------------------------  | 
|
306  | 
NSLIM_add_minus  | 
|
307  | 
----------------------------------------------*)  | 
|
308  | 
Goal "[| f -- x --NS> l; g -- x --NS> m |] \  | 
|
309  | 
\ ==> (%x. f(x) + -g(x)) -- x --NS> (l + -m)";  | 
|
310  | 
by (blast_tac (claset() addDs [NSLIM_add,NSLIM_minus]) 1);  | 
|
311  | 
qed "NSLIM_add_minus";  | 
|
312  | 
||
313  | 
Goal "[| f -- x --> l; g -- x --> m |] \  | 
|
314  | 
\ ==> (%x. f(x) + -g(x)) -- x --> (l + -m)";  | 
|
315  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,  | 
|
316  | 
NSLIM_add_minus]) 1);  | 
|
317  | 
qed "LIM_add_minus2";  | 
|
318  | 
||
319  | 
(*-----------------------------  | 
|
320  | 
NSLIM_inverse  | 
|
321  | 
-----------------------------*)  | 
|
322  | 
Goalw [NSLIM_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
323  | 
"[| f -- a --NS> L; L \\<noteq> 0 |] \  | 
| 10751 | 324  | 
\ ==> (%x. inverse(f(x))) -- a --NS> (inverse L)";  | 
325  | 
by (Clarify_tac 1);  | 
|
326  | 
by (dtac spec 1);  | 
|
327  | 
by (auto_tac (claset(),  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
328  | 
simpset() addsimps [hypreal_of_real_approx_inverse]));  | 
| 10751 | 329  | 
qed "NSLIM_inverse";  | 
330  | 
||
331  | 
Goal "[| f -- a --> L; \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
332  | 
\ L \\<noteq> 0 |] ==> (%x. inverse(f(x))) -- a --> (inverse L)";  | 
| 10751 | 333  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_inverse]) 1);  | 
334  | 
qed "LIM_inverse";  | 
|
335  | 
||
336  | 
(*------------------------------  | 
|
337  | 
NSLIM_zero  | 
|
338  | 
------------------------------*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
339  | 
Goal "f -- a --NS> l ==> (%x. f(x) + -l) -- a --NS> 0";  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
340  | 
by (res_inst_tac [("z1","l")] ((real_add_minus RS subst)) 1);
 | 
| 10751 | 341  | 
by (rtac NSLIM_add_minus 1 THEN Auto_tac);  | 
342  | 
qed "NSLIM_zero";  | 
|
343  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
344  | 
Goal "f -- a --> l ==> (%x. f(x) + -l) -- a --> 0";  | 
| 10751 | 345  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_zero]) 1);  | 
346  | 
qed "LIM_zero2";  | 
|
347  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
348  | 
Goal "(%x. f(x) - l) -- x --NS> 0 ==> f -- x --NS> l";  | 
| 10751 | 349  | 
by (dres_inst_tac [("g","%x. l"),("m","l")] NSLIM_add 1);
 | 
350  | 
by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));  | 
|
351  | 
qed "NSLIM_zero_cancel";  | 
|
352  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
353  | 
Goal "(%x. f(x) - l) -- x --> 0 ==> f -- x --> l";  | 
| 10751 | 354  | 
by (dres_inst_tac [("g","%x. l"),("m","l")] LIM_add 1);
 | 
355  | 
by (auto_tac (claset(),simpset() addsimps [real_diff_def, real_add_assoc]));  | 
|
356  | 
qed "LIM_zero_cancel";  | 
|
357  | 
||
358  | 
||
359  | 
(*--------------------------  | 
|
360  | 
NSLIM_not_zero  | 
|
361  | 
--------------------------*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
362  | 
Goalw [NSLIM_def] "k \\<noteq> 0 ==> ~ ((%x. k) -- x --NS> 0)";  | 
| 10751 | 363  | 
by Auto_tac;  | 
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
364  | 
by (res_inst_tac [("x","hypreal_of_real x + epsilon")] exI 1);
 | 
| 
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
365  | 
by (auto_tac (claset() addIs [Infinitesimal_add_approx_self RS approx_sym],  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
366  | 
simpset() addsimps [hypreal_epsilon_not_zero]));  | 
| 10751 | 367  | 
qed "NSLIM_not_zero";  | 
368  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
369  | 
(* [| k \\<noteq> 0; (%x. k) -- x --NS> 0 |] ==> R *)  | 
| 10751 | 370  | 
bind_thm("NSLIM_not_zeroE", NSLIM_not_zero RS notE);
 | 
371  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
372  | 
Goal "k \\<noteq> 0 ==> ~ ((%x. k) -- x --> 0)";  | 
| 10751 | 373  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_not_zero]) 1);  | 
374  | 
qed "LIM_not_zero2";  | 
|
375  | 
||
376  | 
(*-------------------------------------  | 
|
377  | 
NSLIM of constant function  | 
|
378  | 
-------------------------------------*)  | 
|
379  | 
Goal "(%x. k) -- x --NS> L ==> k = L";  | 
|
380  | 
by (rtac ccontr 1);  | 
|
381  | 
by (dtac NSLIM_zero 1);  | 
|
382  | 
by (rtac NSLIM_not_zeroE 1 THEN assume_tac 2);  | 
|
383  | 
by (arith_tac 1);  | 
|
384  | 
qed "NSLIM_const_eq";  | 
|
385  | 
||
386  | 
Goal "(%x. k) -- x --> L ==> k = L";  | 
|
387  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff,  | 
|
388  | 
NSLIM_const_eq]) 1);  | 
|
389  | 
qed "LIM_const_eq2";  | 
|
390  | 
||
391  | 
(*------------------------  | 
|
392  | 
NS Limit is Unique  | 
|
393  | 
------------------------*)  | 
|
394  | 
(* can actually be proved more easily by unfolding def! *)  | 
|
395  | 
Goal "[| f -- x --NS> L; f -- x --NS> M |] ==> L = M";  | 
|
396  | 
by (dtac NSLIM_minus 1);  | 
|
397  | 
by (dtac NSLIM_add 1 THEN assume_tac 1);  | 
|
398  | 
by (auto_tac (claset() addSDs [NSLIM_const_eq RS sym], simpset()));  | 
|
399  | 
qed "NSLIM_unique";  | 
|
400  | 
||
401  | 
Goal "[| f -- x --> L; f -- x --> M |] ==> L = M";  | 
|
402  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_unique]) 1);  | 
|
403  | 
qed "LIM_unique2";  | 
|
404  | 
||
405  | 
(*--------------------  | 
|
406  | 
NSLIM_mult_zero  | 
|
407  | 
--------------------*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
408  | 
Goal "[| f -- x --NS> 0; g -- x --NS> 0 |] \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
409  | 
\ ==> (%x. f(x)*g(x)) -- x --NS> 0";  | 
| 10751 | 410  | 
by (dtac NSLIM_mult 1 THEN Auto_tac);  | 
411  | 
qed "NSLIM_mult_zero";  | 
|
412  | 
||
413  | 
(* we can use the corresponding thm LIM_mult2 *)  | 
|
414  | 
(* for standard definition of limit *)  | 
|
415  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
416  | 
Goal "[| f -- x --> 0; g -- x --> 0 |] \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
417  | 
\ ==> (%x. f(x)*g(x)) -- x --> 0";  | 
| 10751 | 418  | 
by (dtac LIM_mult2 1 THEN Auto_tac);  | 
419  | 
qed "LIM_mult_zero2";  | 
|
420  | 
||
421  | 
(*----------------------------  | 
|
422  | 
NSLIM_self  | 
|
423  | 
----------------------------*)  | 
|
424  | 
Goalw [NSLIM_def] "(%x. x) -- a --NS> a";  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
425  | 
by (auto_tac (claset() addIs [starfun_Idfun_approx],simpset()));  | 
| 10751 | 426  | 
qed "NSLIM_self";  | 
427  | 
||
428  | 
Goal "(%x. x) -- a --> a";  | 
|
429  | 
by (simp_tac (simpset() addsimps [LIM_NSLIM_iff,NSLIM_self]) 1);  | 
|
430  | 
qed "LIM_self2";  | 
|
431  | 
||
432  | 
(*-----------------------------------------------------------------------------  | 
|
433  | 
Derivatives and Continuity - NS and Standard properties  | 
|
434  | 
-----------------------------------------------------------------------------*)  | 
|
435  | 
(*---------------  | 
|
436  | 
Continuity  | 
|
437  | 
---------------*)  | 
|
438  | 
||
439  | 
Goalw [isNSCont_def]  | 
|
| 11383 | 440  | 
"[| isNSCont f a; y \\<approx> hypreal_of_real a |] \  | 
441  | 
\ ==> (*f* f) y \\<approx> hypreal_of_real (f a)";  | 
|
| 10751 | 442  | 
by (Blast_tac 1);  | 
443  | 
qed "isNSContD";  | 
|
444  | 
||
445  | 
Goalw [isNSCont_def,NSLIM_def]  | 
|
446  | 
"isNSCont f a ==> f -- a --NS> (f a) ";  | 
|
447  | 
by (Blast_tac 1);  | 
|
448  | 
qed "isNSCont_NSLIM";  | 
|
449  | 
||
450  | 
Goalw [isNSCont_def,NSLIM_def]  | 
|
451  | 
"f -- a --NS> (f a) ==> isNSCont f a";  | 
|
452  | 
by Auto_tac;  | 
|
453  | 
by (res_inst_tac [("Q","y = hypreal_of_real a")] 
 | 
|
454  | 
(excluded_middle RS disjE) 1);  | 
|
455  | 
by Auto_tac;  | 
|
456  | 
qed "NSLIM_isNSCont";  | 
|
457  | 
||
458  | 
(*-----------------------------------------------------  | 
|
459  | 
NS continuity can be defined using NS Limit in  | 
|
460  | 
similar fashion to standard def of continuity  | 
|
461  | 
-----------------------------------------------------*)  | 
|
462  | 
Goal "(isNSCont f a) = (f -- a --NS> (f a))";  | 
|
463  | 
by (blast_tac (claset() addIs [isNSCont_NSLIM,NSLIM_isNSCont]) 1);  | 
|
464  | 
qed "isNSCont_NSLIM_iff";  | 
|
465  | 
||
466  | 
(*----------------------------------------------  | 
|
467  | 
Hence, NS continuity can be given  | 
|
468  | 
in terms of standard limit  | 
|
469  | 
---------------------------------------------*)  | 
|
470  | 
Goal "(isNSCont f a) = (f -- a --> (f a))";  | 
|
471  | 
by (asm_full_simp_tac (simpset() addsimps  | 
|
472  | 
[LIM_NSLIM_iff,isNSCont_NSLIM_iff]) 1);  | 
|
473  | 
qed "isNSCont_LIM_iff";  | 
|
474  | 
||
475  | 
(*-----------------------------------------------  | 
|
476  | 
Moreover, it's trivial now that NS continuity  | 
|
477  | 
is equivalent to standard continuity  | 
|
478  | 
-----------------------------------------------*)  | 
|
479  | 
Goalw [isCont_def] "(isNSCont f a) = (isCont f a)";  | 
|
480  | 
by (rtac isNSCont_LIM_iff 1);  | 
|
481  | 
qed "isNSCont_isCont_iff";  | 
|
482  | 
||
483  | 
(*----------------------------------------  | 
|
484  | 
Standard continuity ==> NS continuity  | 
|
485  | 
----------------------------------------*)  | 
|
486  | 
Goal "isCont f a ==> isNSCont f a";  | 
|
487  | 
by (etac (isNSCont_isCont_iff RS iffD2) 1);  | 
|
488  | 
qed "isCont_isNSCont";  | 
|
489  | 
||
490  | 
(*----------------------------------------  | 
|
491  | 
NS continuity ==> Standard continuity  | 
|
492  | 
----------------------------------------*)  | 
|
493  | 
Goal "isNSCont f a ==> isCont f a";  | 
|
494  | 
by (etac (isNSCont_isCont_iff RS iffD1) 1);  | 
|
495  | 
qed "isNSCont_isCont";  | 
|
496  | 
||
497  | 
(*--------------------------------------------------------------------------  | 
|
498  | 
Alternative definition of continuity  | 
|
499  | 
--------------------------------------------------------------------------*)  | 
|
500  | 
(* Prove equivalence between NS limits - *)  | 
|
501  | 
(* seems easier than using standard def *)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
502  | 
Goalw [NSLIM_def] "(f -- a --NS> L) = ((%h. f(a + h)) -- 0 --NS> L)";  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
503  | 
by Auto_tac;  | 
| 10751 | 504  | 
by (dres_inst_tac [("x","hypreal_of_real a + x")] spec 1);
 | 
505  | 
by (dres_inst_tac [("x","-hypreal_of_real a + x")] spec 2);
 | 
|
| 11176 | 506  | 
by Safe_tac;  | 
| 10751 | 507  | 
by (Asm_full_simp_tac 1);  | 
508  | 
by (rtac ((mem_infmal_iff RS iffD2) RS  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
509  | 
(Infinitesimal_add_approx_self RS approx_sym)) 1);  | 
| 
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
510  | 
by (rtac (approx_minus_iff2 RS iffD1) 4);  | 
| 10751 | 511  | 
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 3);  | 
512  | 
by (res_inst_tac [("z","x")] eq_Abs_hypreal 2);
 | 
|
513  | 
by (res_inst_tac [("z","x")] eq_Abs_hypreal 4);
 | 
|
514  | 
by (auto_tac (claset(),  | 
|
515  | 
simpset() addsimps [starfun, hypreal_of_real_def, hypreal_minus,  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
516  | 
hypreal_add, real_add_assoc, approx_refl, hypreal_zero_def]));  | 
| 10751 | 517  | 
qed "NSLIM_h_iff";  | 
518  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
519  | 
Goal "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)";  | 
| 10751 | 520  | 
by (rtac NSLIM_h_iff 1);  | 
521  | 
qed "NSLIM_isCont_iff";  | 
|
522  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
523  | 
Goal "(f -- a --> f a) = ((%h. f(a + h)) -- 0 --> f(a))";  | 
| 10751 | 524  | 
by (simp_tac (simpset() addsimps [LIM_NSLIM_iff, NSLIM_isCont_iff]) 1);  | 
525  | 
qed "LIM_isCont_iff";  | 
|
526  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
527  | 
Goalw [isCont_def] "(isCont f x) = ((%h. f(x + h)) -- 0 --> f(x))";  | 
| 10751 | 528  | 
by (simp_tac (simpset() addsimps [LIM_isCont_iff]) 1);  | 
529  | 
qed "isCont_iff";  | 
|
530  | 
||
531  | 
(*--------------------------------------------------------------------------  | 
|
532  | 
Immediate application of nonstandard criterion for continuity can offer  | 
|
533  | 
very simple proofs of some standard property of continuous functions  | 
|
534  | 
--------------------------------------------------------------------------*)  | 
|
535  | 
(*------------------------  | 
|
536  | 
sum continuous  | 
|
537  | 
------------------------*)  | 
|
538  | 
Goal "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) + g(x)) a";  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
539  | 
by (auto_tac (claset() addIs [approx_add],  | 
| 10751 | 540  | 
simpset() addsimps [isNSCont_isCont_iff RS sym, isNSCont_def]));  | 
541  | 
qed "isCont_add";  | 
|
542  | 
||
543  | 
(*------------------------  | 
|
544  | 
mult continuous  | 
|
545  | 
------------------------*)  | 
|
546  | 
Goal "[| isCont f a; isCont g a |] ==> isCont (%x. f(x) * g(x)) a";  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
547  | 
by (auto_tac (claset() addSIs [starfun_mult_HFinite_approx],  | 
| 10751 | 548  | 
simpset() delsimps [starfun_mult RS sym]  | 
549  | 
addsimps [isNSCont_isCont_iff RS sym, isNSCont_def]));  | 
|
550  | 
qed "isCont_mult";  | 
|
551  | 
||
552  | 
(*-------------------------------------------  | 
|
553  | 
composition of continuous functions  | 
|
554  | 
Note very short straightforard proof!  | 
|
555  | 
------------------------------------------*)  | 
|
556  | 
Goal "[| isCont f a; isCont g (f a) |] \  | 
|
557  | 
\ ==> isCont (g o f) a";  | 
|
558  | 
by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym,  | 
|
559  | 
isNSCont_def,starfun_o RS sym]));  | 
|
560  | 
qed "isCont_o";  | 
|
561  | 
||
562  | 
Goal "[| isCont f a; isCont g (f a) |] \  | 
|
563  | 
\ ==> isCont (%x. g (f x)) a";  | 
|
564  | 
by (auto_tac (claset() addDs [isCont_o],simpset() addsimps [o_def]));  | 
|
565  | 
qed "isCont_o2";  | 
|
566  | 
||
567  | 
Goalw [isNSCont_def] "isNSCont f a ==> isNSCont (%x. - f x) a";  | 
|
568  | 
by Auto_tac;  | 
|
569  | 
qed "isNSCont_minus";  | 
|
570  | 
||
571  | 
Goal "isCont f a ==> isCont (%x. - f x) a";  | 
|
572  | 
by (auto_tac (claset(),simpset() addsimps [isNSCont_isCont_iff RS sym,  | 
|
573  | 
isNSCont_minus]));  | 
|
574  | 
qed "isCont_minus";  | 
|
575  | 
||
576  | 
Goalw [isCont_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
577  | 
"[| isCont f x; f x \\<noteq> 0 |] ==> isCont (%x. inverse (f x)) x";  | 
| 10751 | 578  | 
by (blast_tac (claset() addIs [LIM_inverse]) 1);  | 
579  | 
qed "isCont_inverse";  | 
|
580  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
581  | 
Goal "[| isNSCont f x; f x \\<noteq> 0 |] ==> isNSCont (%x. inverse (f x)) x";  | 
| 10751 | 582  | 
by (auto_tac (claset() addIs [isCont_inverse],simpset() addsimps  | 
583  | 
[isNSCont_isCont_iff]));  | 
|
584  | 
qed "isNSCont_inverse";  | 
|
585  | 
||
586  | 
Goalw [real_diff_def]  | 
|
587  | 
"[| isCont f a; isCont g a |] ==> isCont (%x. f(x) - g(x)) a";  | 
|
588  | 
by (auto_tac (claset() addIs [isCont_add,isCont_minus],simpset()));  | 
|
589  | 
qed "isCont_diff";  | 
|
590  | 
||
591  | 
Goalw [isCont_def] "isCont (%x. k) a";  | 
|
592  | 
by (Simp_tac 1);  | 
|
593  | 
qed "isCont_const";  | 
|
594  | 
Addsimps [isCont_const];  | 
|
595  | 
||
596  | 
Goalw [isNSCont_def] "isNSCont (%x. k) a";  | 
|
597  | 
by (Simp_tac 1);  | 
|
598  | 
qed "isNSCont_const";  | 
|
599  | 
Addsimps [isNSCont_const];  | 
|
600  | 
||
601  | 
Goalw [isNSCont_def] "isNSCont abs a";  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
602  | 
by (auto_tac (claset() addIs [approx_hrabs],  | 
| 10751 | 603  | 
simpset() addsimps [hypreal_of_real_hrabs RS sym,  | 
604  | 
starfun_rabs_hrabs]));  | 
|
605  | 
qed "isNSCont_rabs";  | 
|
606  | 
Addsimps [isNSCont_rabs];  | 
|
607  | 
||
608  | 
Goal "isCont abs a";  | 
|
609  | 
by (auto_tac (claset(), simpset() addsimps [isNSCont_isCont_iff RS sym]));  | 
|
610  | 
qed "isCont_rabs";  | 
|
611  | 
Addsimps [isCont_rabs];  | 
|
612  | 
||
613  | 
(****************************************************************  | 
|
614  | 
(%* Leave as commented until I add topology theory or remove? *%)  | 
|
615  | 
(%*------------------------------------------------------------  | 
|
616  | 
Elementary topology proof for a characterisation of  | 
|
617  | 
continuity now: a function f is continuous if and only  | 
|
| 11383 | 618  | 
  if the inverse image, {x. f(x) \\<in> A}, of any open set A 
 | 
| 10751 | 619  | 
is always an open set  | 
620  | 
------------------------------------------------------------*%)  | 
|
| 11383 | 621  | 
Goal "[| isNSopen A; \\<forall>x. isNSCont f x |] \  | 
622  | 
\              ==> isNSopen {x. f x \\<in> A}";
 | 
|
| 10751 | 623  | 
by (auto_tac (claset(),simpset() addsimps [isNSopen_iff1]));  | 
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
624  | 
by (dtac (mem_monad_approx RS approx_sym) 1);  | 
| 10751 | 625  | 
by (dres_inst_tac [("x","a")] spec 1);
 | 
626  | 
by (dtac isNSContD 1 THEN assume_tac 1);  | 
|
627  | 
by (dtac bspec 1 THEN assume_tac 1);  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
628  | 
by (dres_inst_tac [("x","( *f* f) x")] approx_mem_monad2 1);
 | 
| 10751 | 629  | 
by (blast_tac (claset() addIs [starfun_mem_starset]) 1);  | 
630  | 
qed "isNSCont_isNSopen";  | 
|
631  | 
||
632  | 
Goalw [isNSCont_def]  | 
|
| 11383 | 633  | 
          "\\<forall>A. isNSopen A --> isNSopen {x. f x \\<in> A} \
 | 
| 10751 | 634  | 
\ ==> isNSCont f x";  | 
635  | 
by (auto_tac (claset() addSIs [(mem_infmal_iff RS iffD1) RS  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
636  | 
(approx_minus_iff RS iffD2)],simpset() addsimps  | 
| 10751 | 637  | 
[Infinitesimal_def,SReal_iff]));  | 
638  | 
by (dres_inst_tac [("x","{z. abs(z + -f(x)) < ya}")] spec 1);
 | 
|
639  | 
by (etac (isNSopen_open_interval RSN (2,impE)) 1);  | 
|
640  | 
by (auto_tac (claset(),simpset() addsimps [isNSopen_def,isNSnbhd_def]));  | 
|
641  | 
by (dres_inst_tac [("x","x")] spec 1);
 | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
642  | 
by (auto_tac (claset() addDs [approx_sym RS approx_mem_monad],  | 
| 10751 | 643  | 
simpset() addsimps [hypreal_of_real_zero RS sym,STAR_starfun_rabs_add_minus]));  | 
644  | 
qed "isNSopen_isNSCont";  | 
|
645  | 
||
| 11383 | 646  | 
Goal "(\\<forall>x. isNSCont f x) = \  | 
647  | 
\     (\\<forall>A. isNSopen A --> isNSopen {x. f(x) \\<in> A})";
 | 
|
| 10751 | 648  | 
by (blast_tac (claset() addIs [isNSCont_isNSopen,  | 
649  | 
isNSopen_isNSCont]) 1);  | 
|
650  | 
qed "isNSCont_isNSopen_iff";  | 
|
651  | 
||
652  | 
(%*------- Standard version of same theorem --------*%)  | 
|
| 11383 | 653  | 
Goal "(\\<forall>x. isCont f x) = \  | 
654  | 
\         (\\<forall>A. isopen A --> isopen {x. f(x) \\<in> A})";
 | 
|
| 10751 | 655  | 
by (auto_tac (claset() addSIs [isNSCont_isNSopen_iff],  | 
656  | 
simpset() addsimps [isNSopen_isopen_iff RS sym,  | 
|
657  | 
isNSCont_isCont_iff RS sym]));  | 
|
658  | 
qed "isCont_isopen_iff";  | 
|
659  | 
*******************************************************************)  | 
|
660  | 
||
661  | 
(*-----------------------------------------------------------------  | 
|
662  | 
Uniform continuity  | 
|
663  | 
------------------------------------------------------------------*)  | 
|
664  | 
Goalw [isNSUCont_def]  | 
|
| 11383 | 665  | 
"[| isNSUCont f; x \\<approx> y|] ==> (*f* f) x \\<approx> (*f* f) y";  | 
| 10751 | 666  | 
by (Blast_tac 1);  | 
667  | 
qed "isNSUContD";  | 
|
668  | 
||
669  | 
Goalw [isUCont_def,isCont_def,LIM_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
670  | 
"isUCont f ==> isCont f x";  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
671  | 
by (Clarify_tac 1);  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
672  | 
by (dtac spec 1);  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
673  | 
by (Blast_tac 1);  | 
| 10751 | 674  | 
qed "isUCont_isCont";  | 
675  | 
||
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
676  | 
Goalw [isNSUCont_def,isUCont_def,approx_def]  | 
| 10751 | 677  | 
"isUCont f ==> isNSUCont f";  | 
678  | 
by (asm_full_simp_tac (simpset() addsimps  | 
|
679  | 
[Infinitesimal_FreeUltrafilterNat_iff]) 1);  | 
|
| 11176 | 680  | 
by Safe_tac;  | 
| 10751 | 681  | 
by (res_inst_tac [("z","x")] eq_Abs_hypreal 1);
 | 
682  | 
by (res_inst_tac [("z","y")] eq_Abs_hypreal 1);
 | 
|
683  | 
by (auto_tac (claset(),simpset() addsimps [starfun,  | 
|
684  | 
hypreal_minus, hypreal_add]));  | 
|
685  | 
by (rtac bexI 1 THEN rtac lemma_hyprel_refl 2 THEN Step_tac 1);  | 
|
686  | 
by (dres_inst_tac [("x","u")] spec 1 THEN Clarify_tac 1);
 | 
|
687  | 
by (dres_inst_tac [("x","s")] spec 1 THEN Clarify_tac 1);
 | 
|
| 11383 | 688  | 
by (subgoal_tac "\\<forall>n::nat. abs ((xa n) + - (xb n)) < s --> abs (f (xa n) + - f (xb n)) < u" 1);  | 
| 10751 | 689  | 
by (Blast_tac 2);  | 
| 11383 | 690  | 
by (thin_tac "\\<forall>x y. abs (x + - y) < s --> abs (f x + - f y) < u" 1);  | 
| 10751 | 691  | 
by (dtac FreeUltrafilterNat_all 1);  | 
692  | 
by (Ultra_tac 1);  | 
|
693  | 
qed "isUCont_isNSUCont";  | 
|
694  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
695  | 
Goal "\\<forall>s. 0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \  | 
| 11383 | 696  | 
\ ==> \\<forall>n::nat. \\<exists>z y. \  | 
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
697  | 
\ abs(z + -y) < inverse(real(Suc n)) & \  | 
| 11383 | 698  | 
\ r \\<le> abs(f z + -f y)";  | 
| 
10778
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
699  | 
by (Clarify_tac 1);  | 
| 
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
700  | 
by (cut_inst_tac [("n1","n")]
 | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
701  | 
(real_of_nat_Suc_gt_zero RS real_inverse_gt_0) 1);  | 
| 10751 | 702  | 
by Auto_tac;  | 
703  | 
val lemma_LIMu = result();  | 
|
704  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
705  | 
Goal "\\<forall>s. 0 < s --> (\\<exists>z y. abs (z + - y) < s & r \\<le> abs (f z + -f y)) \  | 
| 11383 | 706  | 
\ ==> \\<exists>X Y. \\<forall>n::nat. \  | 
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
707  | 
\ abs(X n + -(Y n)) < inverse(real(Suc n)) & \  | 
| 11383 | 708  | 
\ r \\<le> abs(f (X n) + -f (Y n))";  | 
| 10751 | 709  | 
by (dtac lemma_LIMu 1);  | 
710  | 
by (dtac choice 1);  | 
|
| 11176 | 711  | 
by Safe_tac;  | 
| 10751 | 712  | 
by (dtac choice 1);  | 
713  | 
by (Blast_tac 1);  | 
|
714  | 
val lemma_skolemize_LIM2u = result();  | 
|
715  | 
||
| 11383 | 716  | 
Goal "\\<forall>n. abs (X n + -Y n) < inverse (real(Suc n)) & \  | 
717  | 
\ r \\<le> abs (f (X n) + - f(Y n)) ==> \  | 
|
718  | 
\ \\<forall>n. abs (X n + - Y n) < inverse (real(Suc n))";  | 
|
| 10751 | 719  | 
by (Auto_tac );  | 
720  | 
val lemma_simpu = result();  | 
|
721  | 
||
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
722  | 
Goalw [isNSUCont_def,isUCont_def,approx_def]  | 
| 10751 | 723  | 
"isNSUCont f ==> isUCont f";  | 
724  | 
by (asm_full_simp_tac (simpset() addsimps  | 
|
725  | 
[Infinitesimal_FreeUltrafilterNat_iff]) 1);  | 
|
726  | 
by (EVERY1[Step_tac, rtac ccontr, Asm_full_simp_tac]);  | 
|
727  | 
by (fold_tac [real_le_def]);  | 
|
728  | 
by (dtac lemma_skolemize_LIM2u 1);  | 
|
| 11176 | 729  | 
by Safe_tac;  | 
| 10834 | 730  | 
by (dres_inst_tac [("x","Abs_hypreal(hyprel``{X})")] spec 1);
 | 
731  | 
by (dres_inst_tac [("x","Abs_hypreal(hyprel``{Y})")] spec 1);
 | 
|
| 10751 | 732  | 
by (asm_full_simp_tac  | 
733  | 
(simpset() addsimps [starfun, hypreal_minus,hypreal_add]) 1);  | 
|
734  | 
by Auto_tac;  | 
|
735  | 
by (dtac (lemma_simpu RS real_seq_to_hypreal_Infinitesimal2) 1);  | 
|
736  | 
by (asm_full_simp_tac (simpset() addsimps  | 
|
737  | 
[Infinitesimal_FreeUltrafilterNat_iff, hypreal_minus,hypreal_add]) 1);  | 
|
738  | 
by (Blast_tac 1);  | 
|
739  | 
by (rotate_tac 2 1);  | 
|
740  | 
by (dres_inst_tac [("x","r")] spec 1);
 | 
|
741  | 
by (Clarify_tac 1);  | 
|
742  | 
by (dtac FreeUltrafilterNat_all 1);  | 
|
743  | 
by (Ultra_tac 1);  | 
|
744  | 
qed "isNSUCont_isUCont";  | 
|
745  | 
||
746  | 
(*------------------------------------------------------------------  | 
|
747  | 
Derivatives  | 
|
748  | 
------------------------------------------------------------------*)  | 
|
749  | 
Goalw [deriv_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
750  | 
"(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --> D)";  | 
| 10751 | 751  | 
by (Blast_tac 1);  | 
752  | 
qed "DERIV_iff";  | 
|
753  | 
||
754  | 
Goalw [deriv_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
755  | 
"(DERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)";  | 
| 10751 | 756  | 
by (simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);  | 
757  | 
qed "DERIV_NS_iff";  | 
|
758  | 
||
759  | 
Goalw [deriv_def]  | 
|
760  | 
"DERIV f x :> D \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
761  | 
\ ==> (%h. (f(x + h) + - f(x))/h) -- 0 --> D";  | 
| 10751 | 762  | 
by (Blast_tac 1);  | 
763  | 
qed "DERIVD";  | 
|
764  | 
||
765  | 
Goalw [deriv_def] "DERIV f x :> D ==> \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
766  | 
\ (%h. (f(x + h) + - f(x))/h) -- 0 --NS> D";  | 
| 10751 | 767  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_NSLIM_iff]) 1);  | 
768  | 
qed "NS_DERIVD";  | 
|
769  | 
||
770  | 
(* Uniqueness *)  | 
|
771  | 
Goalw [deriv_def]  | 
|
772  | 
"[| DERIV f x :> D; DERIV f x :> E |] ==> D = E";  | 
|
773  | 
by (blast_tac (claset() addIs [LIM_unique]) 1);  | 
|
774  | 
qed "DERIV_unique";  | 
|
775  | 
||
776  | 
Goalw [nsderiv_def]  | 
|
777  | 
"[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E";  | 
|
778  | 
by (cut_facts_tac [Infinitesimal_epsilon, hypreal_epsilon_not_zero] 1);  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
779  | 
by (auto_tac (claset() addSDs [inst "x" "epsilon" bspec]  | 
| 10751 | 780  | 
addSIs [inj_hypreal_of_real RS injD]  | 
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
781  | 
addDs [approx_trans3],  | 
| 10751 | 782  | 
simpset()));  | 
783  | 
qed "NSDeriv_unique";  | 
|
784  | 
||
785  | 
(*------------------------------------------------------------------------  | 
|
786  | 
Differentiable  | 
|
787  | 
------------------------------------------------------------------------*)  | 
|
788  | 
||
789  | 
Goalw [differentiable_def]  | 
|
| 11383 | 790  | 
"f differentiable x ==> \\<exists>D. DERIV f x :> D";  | 
| 10751 | 791  | 
by (assume_tac 1);  | 
792  | 
qed "differentiableD";  | 
|
793  | 
||
794  | 
Goalw [differentiable_def]  | 
|
795  | 
"DERIV f x :> D ==> f differentiable x";  | 
|
796  | 
by (Blast_tac 1);  | 
|
797  | 
qed "differentiableI";  | 
|
798  | 
||
799  | 
Goalw [NSdifferentiable_def]  | 
|
| 11383 | 800  | 
"f NSdifferentiable x ==> \\<exists>D. NSDERIV f x :> D";  | 
| 10751 | 801  | 
by (assume_tac 1);  | 
802  | 
qed "NSdifferentiableD";  | 
|
803  | 
||
804  | 
Goalw [NSdifferentiable_def]  | 
|
805  | 
"NSDERIV f x :> D ==> f NSdifferentiable x";  | 
|
806  | 
by (Blast_tac 1);  | 
|
807  | 
qed "NSdifferentiableI";  | 
|
808  | 
||
809  | 
(*--------------------------------------------------------  | 
|
810  | 
Alternative definition for differentiability  | 
|
811  | 
-------------------------------------------------------*)  | 
|
812  | 
||
813  | 
Goalw [LIM_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
814  | 
"((%h. (f(a + h) + - f(a))/h) -- 0 --> D) = \  | 
| 10751 | 815  | 
\ ((%x. (f(x) + -f(a)) / (x + -a)) -- a --> D)";  | 
| 11176 | 816  | 
by Safe_tac;  | 
| 10751 | 817  | 
by (ALLGOALS(dtac spec));  | 
| 11176 | 818  | 
by Safe_tac;  | 
| 10751 | 819  | 
by (Blast_tac 1 THEN Blast_tac 2);  | 
820  | 
by (ALLGOALS(res_inst_tac [("x","s")] exI));
 | 
|
| 11176 | 821  | 
by Safe_tac;  | 
| 10751 | 822  | 
by (dres_inst_tac [("x","x + -a")] spec 1);
 | 
823  | 
by (dres_inst_tac [("x","x + a")] spec 2);
 | 
|
824  | 
by (auto_tac (claset(), simpset() addsimps real_add_ac));  | 
|
825  | 
qed "DERIV_LIM_iff";  | 
|
826  | 
||
827  | 
Goalw [deriv_def] "(DERIV f x :> D) = \  | 
|
828  | 
\ ((%z. (f(z) + -f(x)) / (z + -x)) -- x --> D)";  | 
|
829  | 
by (simp_tac (simpset() addsimps [DERIV_LIM_iff]) 1);  | 
|
830  | 
qed "DERIV_iff2";  | 
|
831  | 
||
832  | 
(*--------------------------------------------------------  | 
|
833  | 
Equivalence of NS and standard defs of differentiation  | 
|
834  | 
-------------------------------------------------------*)  | 
|
835  | 
(*-------------------------------------------  | 
|
836  | 
First NSDERIV in terms of NSLIM  | 
|
837  | 
-------------------------------------------*)  | 
|
838  | 
||
839  | 
(*--- first equivalence ---*)  | 
|
840  | 
Goalw [nsderiv_def,NSLIM_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
841  | 
"(NSDERIV f x :> D) = ((%h. (f(x + h) + - f(x))/h) -- 0 --NS> D)";  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
842  | 
by Auto_tac;  | 
| 10751 | 843  | 
by (dres_inst_tac [("x","xa")] bspec 1);
 | 
844  | 
by (rtac ccontr 3);  | 
|
845  | 
by (dres_inst_tac [("x","h")] spec 3);
 | 
|
846  | 
by (auto_tac (claset(),  | 
|
847  | 
simpset() addsimps [mem_infmal_iff, starfun_lambda_cancel]));  | 
|
848  | 
qed "NSDERIV_NSLIM_iff";  | 
|
849  | 
||
850  | 
(*--- second equivalence ---*)  | 
|
851  | 
Goal "(NSDERIV f x :> D) = \  | 
|
852  | 
\ ((%z. (f(z) + -f(x)) / (z + -x)) -- x --NS> D)";  | 
|
853  | 
by (full_simp_tac (simpset() addsimps  | 
|
854  | 
[NSDERIV_NSLIM_iff, DERIV_LIM_iff, LIM_NSLIM_iff RS sym]) 1);  | 
|
855  | 
qed "NSDERIV_NSLIM_iff2";  | 
|
856  | 
||
857  | 
(* while we're at it! *)  | 
|
858  | 
Goalw [real_diff_def]  | 
|
859  | 
"(NSDERIV f x :> D) = \  | 
|
| 11383 | 860  | 
\ (\\<forall>xa. \  | 
861  | 
\ xa \\<noteq> hypreal_of_real x & xa \\<approx> hypreal_of_real x --> \  | 
|
862  | 
\ (*f* (%z. (f z - f x) / (z - x))) xa \\<approx> hypreal_of_real D)";  | 
|
| 10751 | 863  | 
by (auto_tac (claset(), simpset() addsimps [NSDERIV_NSLIM_iff2, NSLIM_def]));  | 
864  | 
qed "NSDERIV_iff2";  | 
|
865  | 
||
866  | 
||
867  | 
Goal "(NSDERIV f x :> D) ==> \  | 
|
| 11383 | 868  | 
\ (\\<forall>u. \  | 
869  | 
\ u \\<approx> hypreal_of_real x --> \  | 
|
870  | 
\ (*f* (%z. f z - f x)) u \\<approx> hypreal_of_real D * (u - hypreal_of_real x))";  | 
|
| 10751 | 871  | 
by (auto_tac (claset(), simpset() addsimps [NSDERIV_iff2]));  | 
872  | 
by (case_tac "u = hypreal_of_real x" 1);  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
873  | 
by (auto_tac (claset(), simpset() addsimps [hypreal_diff_def]));  | 
| 10751 | 874  | 
by (dres_inst_tac [("x","u")] spec 1);
 | 
875  | 
by Auto_tac;  | 
|
876  | 
by (dres_inst_tac [("c","u - hypreal_of_real x"),("b","hypreal_of_real D")]
 | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
877  | 
approx_mult1 1);  | 
| 10751 | 878  | 
by (ALLGOALS(dtac (hypreal_not_eq_minus_iff RS iffD1)));  | 
| 11383 | 879  | 
by (subgoal_tac "(*f* (%z. z - x)) u \\<noteq> (0::hypreal)" 2);  | 
| 10751 | 880  | 
by (rotate_tac ~1 2);  | 
881  | 
by (auto_tac (claset(),  | 
|
882  | 
simpset() addsimps [real_diff_def, hypreal_diff_def,  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
883  | 
(approx_minus_iff RS iffD1) RS (mem_infmal_iff RS iffD2),  | 
| 10751 | 884  | 
Infinitesimal_subset_HFinite RS subsetD]));  | 
885  | 
qed "NSDERIVD5";  | 
|
886  | 
||
887  | 
Goal "(NSDERIV f x :> D) ==> \  | 
|
| 11383 | 888  | 
\ (\\<forall>h \\<in> Infinitesimal. \  | 
| 10751 | 889  | 
\ ((*f* f)(hypreal_of_real x + h) - \  | 
| 11383 | 890  | 
\ hypreal_of_real (f x))\\<approx> (hypreal_of_real D) * h)";  | 
| 10751 | 891  | 
by (auto_tac (claset(),simpset() addsimps [nsderiv_def]));  | 
892  | 
by (case_tac "h = (0::hypreal)" 1);  | 
|
893  | 
by (auto_tac (claset(),simpset() addsimps [hypreal_diff_def]));  | 
|
894  | 
by (dres_inst_tac [("x","h")] bspec 1);
 | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
895  | 
by (dres_inst_tac [("c","h")] approx_mult1 2);
 | 
| 10751 | 896  | 
by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],  | 
897  | 
simpset() addsimps [hypreal_diff_def]));  | 
|
898  | 
qed "NSDERIVD4";  | 
|
899  | 
||
900  | 
Goal "(NSDERIV f x :> D) ==> \  | 
|
| 11383 | 901  | 
\     (\\<forall>h \\<in> Infinitesimal - {0}. \
 | 
| 10751 | 902  | 
\ ((*f* f)(hypreal_of_real x + h) - \  | 
| 11383 | 903  | 
\ hypreal_of_real (f x))\\<approx> (hypreal_of_real D) * h)";  | 
| 10751 | 904  | 
by (auto_tac (claset(),simpset() addsimps [nsderiv_def]));  | 
905  | 
by (rtac ccontr 1 THEN dres_inst_tac [("x","h")] bspec 1);
 | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
906  | 
by (dres_inst_tac [("c","h")] approx_mult1 2);
 | 
| 10751 | 907  | 
by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite RS subsetD],  | 
908  | 
simpset() addsimps [hypreal_mult_assoc, hypreal_diff_def]));  | 
|
909  | 
qed "NSDERIVD3";  | 
|
910  | 
||
911  | 
(*--------------------------------------------------------------  | 
|
912  | 
Now equivalence between NSDERIV and DERIV  | 
|
913  | 
-------------------------------------------------------------*)  | 
|
914  | 
Goalw [deriv_def] "(NSDERIV f x :> D) = (DERIV f x :> D)";  | 
|
915  | 
by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,LIM_NSLIM_iff]) 1);  | 
|
916  | 
qed "NSDERIV_DERIV_iff";  | 
|
917  | 
||
918  | 
(*---------------------------------------------------  | 
|
919  | 
Differentiability implies continuity  | 
|
920  | 
nice and simple "algebraic" proof  | 
|
921  | 
--------------------------------------------------*)  | 
|
922  | 
Goalw [nsderiv_def]  | 
|
923  | 
"NSDERIV f x :> D ==> isNSCont f x";  | 
|
924  | 
by (auto_tac (claset(),simpset() addsimps  | 
|
925  | 
[isNSCont_NSLIM_iff,NSLIM_def]));  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
926  | 
by (dtac (approx_minus_iff RS iffD1) 1);  | 
| 10751 | 927  | 
by (dtac (hypreal_not_eq_minus_iff RS iffD1) 1);  | 
928  | 
by (dres_inst_tac [("x","-hypreal_of_real x + xa")] bspec 1);
 | 
|
929  | 
by (asm_full_simp_tac (simpset() addsimps  | 
|
930  | 
[hypreal_add_assoc RS sym]) 2);  | 
|
931  | 
by (auto_tac (claset(),simpset() addsimps  | 
|
932  | 
[mem_infmal_iff RS sym,hypreal_add_commute]));  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
933  | 
by (dres_inst_tac [("c","xa + -hypreal_of_real x")] approx_mult1 1);
 | 
| 10751 | 934  | 
by (auto_tac (claset() addIs [Infinitesimal_subset_HFinite  | 
935  | 
RS subsetD],simpset() addsimps [hypreal_mult_assoc]));  | 
|
936  | 
by (dres_inst_tac [("x3","D")] (HFinite_hypreal_of_real RSN
 | 
|
937  | 
(2,Infinitesimal_HFinite_mult) RS (mem_infmal_iff RS iffD1)) 1);  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
938  | 
by (blast_tac (claset() addIs [approx_trans,  | 
| 10751 | 939  | 
hypreal_mult_commute RS subst,  | 
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
940  | 
(approx_minus_iff RS iffD2)]) 1);  | 
| 10751 | 941  | 
qed "NSDERIV_isNSCont";  | 
942  | 
||
943  | 
(* Now Sandard proof *)  | 
|
944  | 
Goal "DERIV f x :> D ==> isCont f x";  | 
|
945  | 
by (asm_full_simp_tac (simpset() addsimps  | 
|
946  | 
[NSDERIV_DERIV_iff RS sym, isNSCont_isCont_iff RS sym,  | 
|
947  | 
NSDERIV_isNSCont]) 1);  | 
|
948  | 
qed "DERIV_isCont";  | 
|
949  | 
||
950  | 
(*----------------------------------------------------------------------------  | 
|
951  | 
Differentiation rules for combinations of functions  | 
|
952  | 
follow from clear, straightforard, algebraic  | 
|
953  | 
manipulations  | 
|
954  | 
----------------------------------------------------------------------------*)  | 
|
955  | 
(*-------------------------  | 
|
956  | 
Constant function  | 
|
957  | 
------------------------*)  | 
|
958  | 
||
959  | 
(* use simple constant nslimit theorem *)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
960  | 
Goal "(NSDERIV (%x. k) x :> 0)";  | 
| 10751 | 961  | 
by (simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1);  | 
962  | 
qed "NSDERIV_const";  | 
|
963  | 
Addsimps [NSDERIV_const];  | 
|
964  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
965  | 
Goal "(DERIV (%x. k) x :> 0)";  | 
| 10751 | 966  | 
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);  | 
967  | 
qed "DERIV_const";  | 
|
968  | 
Addsimps [DERIV_const];  | 
|
969  | 
||
970  | 
(*-----------------------------------------------------  | 
|
971  | 
Sum of functions- proved easily  | 
|
972  | 
----------------------------------------------------*)  | 
|
973  | 
||
974  | 
||
975  | 
Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \  | 
|
976  | 
\ ==> NSDERIV (%x. f x + g x) x :> Da + Db";  | 
|
977  | 
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,  | 
|
| 11176 | 978  | 
NSLIM_def]) 1 THEN REPEAT (Step_tac 1));  | 
| 10751 | 979  | 
by (auto_tac (claset(),  | 
980  | 
simpset() addsimps [hypreal_add_divide_distrib]));  | 
|
981  | 
by (dres_inst_tac [("b","hypreal_of_real Da"),
 | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
982  | 
                   ("d","hypreal_of_real Db")] approx_add 1);
 | 
| 10751 | 983  | 
by (auto_tac (claset(), simpset() addsimps hypreal_add_ac));  | 
984  | 
qed "NSDERIV_add";  | 
|
985  | 
||
986  | 
(* Standard theorem *)  | 
|
987  | 
Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \  | 
|
988  | 
\ ==> DERIV (%x. f x + g x) x :> Da + Db";  | 
|
989  | 
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_add,  | 
|
990  | 
NSDERIV_DERIV_iff RS sym]) 1);  | 
|
991  | 
qed "DERIV_add";  | 
|
992  | 
||
993  | 
(*-----------------------------------------------------  | 
|
994  | 
Product of functions - Proof is trivial but tedious  | 
|
995  | 
and long due to rearrangement of terms  | 
|
996  | 
----------------------------------------------------*)  | 
|
997  | 
||
998  | 
Goal "((a::hypreal)*b) + -(c*d) = (b*(a + -c)) + (c*(b + -d))";  | 
|
999  | 
by (simp_tac (simpset() addsimps [hypreal_add_mult_distrib2]) 1);  | 
|
1000  | 
val lemma_nsderiv1 = result();  | 
|
1001  | 
||
| 11383 | 1002  | 
Goal "[| (x + y) / z = hypreal_of_real D + yb; z \\<noteq> 0; \  | 
1003  | 
\ z \\<in> Infinitesimal; yb \\<in> Infinitesimal |] \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1004  | 
\ ==> x + y \\<approx> 0";  | 
| 10751 | 1005  | 
by (forw_inst_tac [("c1","z")] (hypreal_mult_right_cancel RS iffD2) 1 
 | 
1006  | 
THEN assume_tac 1);  | 
|
1007  | 
by (thin_tac "(x + y) / z = hypreal_of_real D + yb" 1);  | 
|
1008  | 
by (auto_tac (claset() addSIs [Infinitesimal_HFinite_mult2, HFinite_add],  | 
|
1009  | 
simpset() addsimps [hypreal_mult_assoc, mem_infmal_iff RS sym]));  | 
|
1010  | 
by (etac (Infinitesimal_subset_HFinite RS subsetD) 1);  | 
|
1011  | 
val lemma_nsderiv2 = result();  | 
|
1012  | 
||
1013  | 
||
1014  | 
Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \  | 
|
1015  | 
\ ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";  | 
|
| 11176 | 1016  | 
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def]) 1);  | 
1017  | 
by (REPEAT (Step_tac 1));  | 
|
| 10751 | 1018  | 
by (auto_tac (claset(),  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1019  | 
simpset() addsimps [starfun_lambda_cancel, lemma_nsderiv1]));  | 
| 10751 | 1020  | 
by (simp_tac (simpset() addsimps [hypreal_add_divide_distrib]) 1);  | 
1021  | 
by (REPEAT(dtac (bex_Infinitesimal_iff2 RS iffD2) 1));  | 
|
1022  | 
by (auto_tac (claset(),  | 
|
1023  | 
simpset() delsimps [hypreal_times_divide1_eq]  | 
|
1024  | 
addsimps [hypreal_times_divide1_eq RS sym]));  | 
|
1025  | 
by (dres_inst_tac [("D","Db")] lemma_nsderiv2 1);
 | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1026  | 
by (dtac (approx_minus_iff RS iffD2 RS (bex_Infinitesimal_iff2 RS iffD2)) 4);  | 
| 
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1027  | 
by (auto_tac (claset() addSIs [approx_add_mono1],  | 
| 10751 | 1028  | 
simpset() addsimps [hypreal_add_mult_distrib, hypreal_add_mult_distrib2,  | 
1029  | 
hypreal_mult_commute, hypreal_add_assoc]));  | 
|
1030  | 
by (res_inst_tac [("w1","hypreal_of_real Db * hypreal_of_real (f x)")]
 | 
|
1031  | 
(hypreal_add_commute RS subst) 1);  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1032  | 
by (auto_tac (claset() addSIs [Infinitesimal_add_approx_self2 RS approx_sym,  | 
| 10751 | 1033  | 
Infinitesimal_add, Infinitesimal_mult,  | 
1034  | 
Infinitesimal_hypreal_of_real_mult,  | 
|
1035  | 
Infinitesimal_hypreal_of_real_mult2],  | 
|
1036  | 
simpset() addsimps [hypreal_add_assoc RS sym]));  | 
|
1037  | 
qed "NSDERIV_mult";  | 
|
1038  | 
||
1039  | 
Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \  | 
|
1040  | 
\ ==> DERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))";  | 
|
1041  | 
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_mult,  | 
|
1042  | 
NSDERIV_DERIV_iff RS sym]) 1);  | 
|
1043  | 
qed "DERIV_mult";  | 
|
1044  | 
||
1045  | 
(*----------------------------  | 
|
1046  | 
Multiplying by a constant  | 
|
1047  | 
---------------------------*)  | 
|
1048  | 
Goal "NSDERIV f x :> D \  | 
|
1049  | 
\ ==> NSDERIV (%x. c * f x) x :> c*D";  | 
|
1050  | 
by (asm_full_simp_tac  | 
|
1051  | 
(simpset() addsimps [real_times_divide1_eq RS sym, NSDERIV_NSLIM_iff,  | 
|
1052  | 
real_minus_mult_eq2, real_add_mult_distrib2 RS sym]  | 
|
| 
12481
 
ea5d6da573c5
mods due to reorienting and renaming of real_minus_mult_eq1/2
 
nipkow 
parents: 
12018 
diff
changeset
 | 
1053  | 
delsimps [real_times_divide1_eq, real_mult_minus_eq2]) 1);  | 
| 10751 | 1054  | 
by (etac (NSLIM_const RS NSLIM_mult) 1);  | 
1055  | 
qed "NSDERIV_cmult";  | 
|
1056  | 
||
1057  | 
(* let's do the standard proof though theorem *)  | 
|
1058  | 
(* LIM_mult2 follows from a NS proof *)  | 
|
1059  | 
||
1060  | 
Goalw [deriv_def]  | 
|
1061  | 
"DERIV f x :> D \  | 
|
1062  | 
\ ==> DERIV (%x. c * f x) x :> c*D";  | 
|
1063  | 
by (asm_full_simp_tac  | 
|
1064  | 
(simpset() addsimps [real_times_divide1_eq RS sym, NSDERIV_NSLIM_iff,  | 
|
1065  | 
real_minus_mult_eq2, real_add_mult_distrib2 RS sym]  | 
|
| 
12481
 
ea5d6da573c5
mods due to reorienting and renaming of real_minus_mult_eq1/2
 
nipkow 
parents: 
12018 
diff
changeset
 | 
1066  | 
delsimps [real_times_divide1_eq, real_mult_minus_eq2]) 1);  | 
| 10751 | 1067  | 
by (etac (LIM_const RS LIM_mult2) 1);  | 
1068  | 
qed "DERIV_cmult";  | 
|
1069  | 
||
1070  | 
(*--------------------------------  | 
|
1071  | 
Negation of function  | 
|
1072  | 
-------------------------------*)  | 
|
1073  | 
Goal "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D";  | 
|
1074  | 
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff]) 1);  | 
|
1075  | 
by (res_inst_tac [("t","f x")] (real_minus_minus RS subst) 1);
 | 
|
1076  | 
by (asm_simp_tac (simpset() addsimps [real_minus_add_distrib RS sym,  | 
|
| 
12481
 
ea5d6da573c5
mods due to reorienting and renaming of real_minus_mult_eq1/2
 
nipkow 
parents: 
12018 
diff
changeset
 | 
1077  | 
real_mult_minus_eq1]  | 
| 10751 | 1078  | 
delsimps [real_minus_add_distrib, real_minus_minus]) 1);  | 
1079  | 
by (etac NSLIM_minus 1);  | 
|
1080  | 
qed "NSDERIV_minus";  | 
|
1081  | 
||
1082  | 
Goal "DERIV f x :> D \  | 
|
1083  | 
\ ==> DERIV (%x. -(f x)) x :> -D";  | 
|
1084  | 
by (asm_full_simp_tac (simpset() addsimps  | 
|
1085  | 
[NSDERIV_minus,NSDERIV_DERIV_iff RS sym]) 1);  | 
|
1086  | 
qed "DERIV_minus";  | 
|
1087  | 
||
1088  | 
(*-------------------------------  | 
|
1089  | 
Subtraction  | 
|
1090  | 
------------------------------*)  | 
|
1091  | 
Goal "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \  | 
|
1092  | 
\ ==> NSDERIV (%x. f x + -g x) x :> Da + -Db";  | 
|
1093  | 
by (blast_tac (claset() addDs [NSDERIV_add,NSDERIV_minus]) 1);  | 
|
1094  | 
qed "NSDERIV_add_minus";  | 
|
1095  | 
||
1096  | 
Goal "[| DERIV f x :> Da; DERIV g x :> Db |] \  | 
|
1097  | 
\ ==> DERIV (%x. f x + -g x) x :> Da + -Db";  | 
|
1098  | 
by (blast_tac (claset() addDs [DERIV_add,DERIV_minus]) 1);  | 
|
1099  | 
qed "DERIV_add_minus";  | 
|
1100  | 
||
1101  | 
Goalw [real_diff_def]  | 
|
1102  | 
"[| NSDERIV f x :> Da; NSDERIV g x :> Db |] \  | 
|
1103  | 
\ ==> NSDERIV (%x. f x - g x) x :> Da - Db";  | 
|
1104  | 
by (blast_tac (claset() addIs [NSDERIV_add_minus]) 1);  | 
|
1105  | 
qed "NSDERIV_diff";  | 
|
1106  | 
||
1107  | 
Goalw [real_diff_def]  | 
|
1108  | 
"[| DERIV f x :> Da; DERIV g x :> Db |] \  | 
|
1109  | 
\ ==> DERIV (%x. f x - g x) x :> Da - Db";  | 
|
1110  | 
by (blast_tac (claset() addIs [DERIV_add_minus]) 1);  | 
|
1111  | 
qed "DERIV_diff";  | 
|
1112  | 
||
1113  | 
(*---------------------------------------------------------------  | 
|
1114  | 
(NS) Increment  | 
|
1115  | 
---------------------------------------------------------------*)  | 
|
1116  | 
Goalw [increment_def]  | 
|
1117  | 
"f NSdifferentiable x ==> \  | 
|
1118  | 
\ increment f x h = (*f* f) (hypreal_of_real(x) + h) + \  | 
|
1119  | 
\ -hypreal_of_real (f x)";  | 
|
1120  | 
by (Blast_tac 1);  | 
|
1121  | 
qed "incrementI";  | 
|
1122  | 
||
1123  | 
Goal "NSDERIV f x :> D ==> \  | 
|
1124  | 
\ increment f x h = (*f* f) (hypreal_of_real(x) + h) + \  | 
|
1125  | 
\ -hypreal_of_real (f x)";  | 
|
1126  | 
by (etac (NSdifferentiableI RS incrementI) 1);  | 
|
1127  | 
qed "incrementI2";  | 
|
1128  | 
||
1129  | 
(* The Increment theorem -- Keisler p. 65 *)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1130  | 
Goal "[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> 0 |] \  | 
| 11383 | 1131  | 
\ ==> \\<exists>e \\<in> Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h";  | 
| 10751 | 1132  | 
by (forw_inst_tac [("h","h")] incrementI2 1 THEN rewtac nsderiv_def);
 | 
1133  | 
by (dtac bspec 1 THEN Auto_tac);  | 
|
1134  | 
by (dtac (bex_Infinitesimal_iff2 RS iffD2) 1 THEN Step_tac 1);  | 
|
1135  | 
by (forw_inst_tac [("b1","hypreal_of_real(D) + y")] 
 | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1136  | 
((hypreal_mult_right_cancel RS iffD2)) 1);  | 
| 10751 | 1137  | 
by (thin_tac "((*f* f) (hypreal_of_real(x) + h) + \  | 
1138  | 
\ - hypreal_of_real (f x)) / h = hypreal_of_real(D) + y" 2);  | 
|
1139  | 
by (assume_tac 1);  | 
|
1140  | 
by (asm_full_simp_tac (simpset() addsimps [hypreal_times_divide1_eq RS sym]  | 
|
1141  | 
delsimps [hypreal_times_divide1_eq]) 1);  | 
|
1142  | 
by (auto_tac (claset(),  | 
|
1143  | 
simpset() addsimps [hypreal_add_mult_distrib]));  | 
|
1144  | 
qed "increment_thm";  | 
|
1145  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1146  | 
Goal "[| NSDERIV f x :> D; h \\<approx> 0; h \\<noteq> 0 |] \  | 
| 11383 | 1147  | 
\ ==> \\<exists>e \\<in> Infinitesimal. increment f x h = \  | 
| 10751 | 1148  | 
\ hypreal_of_real(D)*h + e*h";  | 
1149  | 
by (blast_tac (claset() addSDs [mem_infmal_iff RS iffD2]  | 
|
1150  | 
addSIs [increment_thm]) 1);  | 
|
1151  | 
qed "increment_thm2";  | 
|
1152  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1153  | 
Goal "[| NSDERIV f x :> D; h \\<approx> 0; h \\<noteq> 0 |] \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1154  | 
\ ==> increment f x h \\<approx> 0";  | 
| 10751 | 1155  | 
by (dtac increment_thm2 1 THEN auto_tac (claset() addSIs  | 
1156  | 
[Infinitesimal_HFinite_mult2,HFinite_add],simpset() addsimps  | 
|
1157  | 
[hypreal_add_mult_distrib RS sym,mem_infmal_iff RS sym]));  | 
|
1158  | 
by (etac (Infinitesimal_subset_HFinite RS subsetD) 1);  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1159  | 
qed "increment_approx_zero";  | 
| 10751 | 1160  | 
|
1161  | 
(*---------------------------------------------------------------  | 
|
1162  | 
Similarly to the above, the chain rule admits an entirely  | 
|
1163  | 
straightforward derivation. Compare this with Harrison's  | 
|
1164  | 
HOL proof of the chain rule, which proved to be trickier and  | 
|
1165  | 
required an alternative characterisation of differentiability-  | 
|
1166  | 
the so-called Carathedory derivative. Our main problem is  | 
|
1167  | 
manipulation of terms.  | 
|
1168  | 
--------------------------------------------------------------*)  | 
|
1169  | 
||
1170  | 
(* lemmas *)  | 
|
1171  | 
Goalw [nsderiv_def]  | 
|
1172  | 
"[| NSDERIV g x :> D; \  | 
|
1173  | 
\ (*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real(g x);\  | 
|
| 11383 | 1174  | 
\ xa \\<in> Infinitesimal;\  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1175  | 
\ xa \\<noteq> 0 \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1176  | 
\ |] ==> D = 0";  | 
| 10751 | 1177  | 
by (dtac bspec 1);  | 
1178  | 
by Auto_tac;  | 
|
1179  | 
qed "NSDERIV_zero";  | 
|
1180  | 
||
1181  | 
(* can be proved differently using NSLIM_isCont_iff *)  | 
|
1182  | 
Goalw [nsderiv_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1183  | 
"[| NSDERIV f x :> D; h \\<in> Infinitesimal; h \\<noteq> 0 |] \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1184  | 
\ ==> (*f* f) (hypreal_of_real(x) + h) + -hypreal_of_real(f x) \\<approx> 0";  | 
| 10751 | 1185  | 
by (asm_full_simp_tac (simpset() addsimps  | 
1186  | 
[mem_infmal_iff RS sym]) 1);  | 
|
1187  | 
by (rtac Infinitesimal_ratio 1);  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1188  | 
by (rtac approx_hypreal_of_real_HFinite 3);  | 
| 10751 | 1189  | 
by Auto_tac;  | 
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1190  | 
qed "NSDERIV_approx";  | 
| 10751 | 1191  | 
|
1192  | 
(*---------------------------------------------------------------  | 
|
1193  | 
from one version of differentiability  | 
|
1194  | 
||
1195  | 
f(x) - f(a)  | 
|
| 11383 | 1196  | 
--------------- \\<approx> Db  | 
| 10751 | 1197  | 
x - a  | 
1198  | 
---------------------------------------------------------------*)  | 
|
1199  | 
Goal "[| NSDERIV f (g x) :> Da; \  | 
|
| 11383 | 1200  | 
\ (*f* g) (hypreal_of_real(x) + xa) \\<noteq> hypreal_of_real (g x); \  | 
1201  | 
\ (*f* g) (hypreal_of_real(x) + xa) \\<approx> hypreal_of_real (g x) \  | 
|
| 10751 | 1202  | 
\ |] ==> ((*f* f) ((*f* g) (hypreal_of_real(x) + xa)) \  | 
1203  | 
\ + - hypreal_of_real (f (g x))) \  | 
|
1204  | 
\ / ((*f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real (g x)) \  | 
|
| 11383 | 1205  | 
\ \\<approx> hypreal_of_real(Da)";  | 
| 10751 | 1206  | 
by (auto_tac (claset(),  | 
1207  | 
simpset() addsimps [NSDERIV_NSLIM_iff2, NSLIM_def]));  | 
|
1208  | 
qed "NSDERIVD1";  | 
|
1209  | 
||
1210  | 
(*--------------------------------------------------------------  | 
|
1211  | 
from other version of differentiability  | 
|
1212  | 
||
1213  | 
f(x + h) - f(x)  | 
|
| 11383 | 1214  | 
----------------- \\<approx> Db  | 
| 10751 | 1215  | 
h  | 
1216  | 
--------------------------------------------------------------*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1217  | 
Goal "[| NSDERIV g x :> Db; xa \\<in> Infinitesimal; xa \\<noteq> 0 |] \  | 
| 10751 | 1218  | 
\ ==> ((*f* g) (hypreal_of_real(x) + xa) + - hypreal_of_real(g x)) / xa \  | 
| 11383 | 1219  | 
\ \\<approx> hypreal_of_real(Db)";  | 
| 10751 | 1220  | 
by (auto_tac (claset(),  | 
1221  | 
simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def,  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1222  | 
mem_infmal_iff, starfun_lambda_cancel]));  | 
| 10751 | 1223  | 
qed "NSDERIVD2";  | 
1224  | 
||
| 11383 | 1225  | 
Goal "(z::hypreal) \\<noteq> 0 ==> x*y = (x*inverse(z))*(z*y)";  | 
| 10751 | 1226  | 
by Auto_tac;  | 
1227  | 
qed "lemma_chain";  | 
|
1228  | 
||
1229  | 
(*------------------------------------------------------  | 
|
1230  | 
This proof uses both definitions of differentiability.  | 
|
1231  | 
------------------------------------------------------*)  | 
|
1232  | 
Goal "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] \  | 
|
1233  | 
\ ==> NSDERIV (f o g) x :> Da * Db";  | 
|
1234  | 
by (asm_simp_tac (simpset() addsimps [NSDERIV_NSLIM_iff,  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1235  | 
NSLIM_def,mem_infmal_iff RS sym]) 1 THEN Step_tac 1);  | 
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1236  | 
by (forw_inst_tac [("f","g")] NSDERIV_approx 1);
 | 
| 10751 | 1237  | 
by (auto_tac (claset(),  | 
1238  | 
simpset() addsimps [starfun_lambda_cancel2, starfun_o RS sym]));  | 
|
1239  | 
by (case_tac "(*f* g) (hypreal_of_real(x) + xa) = hypreal_of_real (g x)" 1);  | 
|
1240  | 
by (dres_inst_tac [("g","g")] NSDERIV_zero 1);
 | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1241  | 
by (auto_tac (claset(), simpset() addsimps [hypreal_divide_def]));  | 
| 10751 | 1242  | 
by (res_inst_tac [("z1","(*f* g) (hypreal_of_real(x) + xa) + -hypreal_of_real (g x)"),
 | 
1243  | 
    ("y1","inverse xa")] (lemma_chain RS ssubst) 1);
 | 
|
1244  | 
by (etac (hypreal_not_eq_minus_iff RS iffD1) 1);  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1245  | 
by (rtac approx_mult_hypreal_of_real 1);  | 
| 10751 | 1246  | 
by (fold_tac [hypreal_divide_def]);  | 
1247  | 
by (blast_tac (claset() addIs [NSDERIVD1,  | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1248  | 
approx_minus_iff RS iffD2]) 1);  | 
| 10751 | 1249  | 
by (blast_tac (claset() addIs [NSDERIVD2]) 1);  | 
1250  | 
qed "NSDERIV_chain";  | 
|
1251  | 
||
1252  | 
(* standard version *)  | 
|
1253  | 
Goal "[| DERIV f (g x) :> Da; \  | 
|
1254  | 
\ DERIV g x :> Db \  | 
|
1255  | 
\ |] ==> DERIV (f o g) x :> Da * Db";  | 
|
1256  | 
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym,  | 
|
1257  | 
NSDERIV_chain]) 1);  | 
|
1258  | 
qed "DERIV_chain";  | 
|
1259  | 
||
1260  | 
Goal "[| DERIV f (g x) :> Da; DERIV g x :> Db |] \  | 
|
1261  | 
\ ==> DERIV (%x. f (g x)) x :> Da * Db";  | 
|
1262  | 
by (auto_tac (claset() addDs [DERIV_chain], simpset() addsimps [o_def]));  | 
|
1263  | 
qed "DERIV_chain2";  | 
|
1264  | 
||
1265  | 
(*------------------------------------------------------------------  | 
|
1266  | 
Differentiation of natural number powers  | 
|
1267  | 
------------------------------------------------------------------*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1268  | 
Goal "NSDERIV (%x. x) x :> 1";  | 
| 10751 | 1269  | 
by (auto_tac (claset(),  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1270  | 
simpset() addsimps [NSDERIV_NSLIM_iff, NSLIM_def ,starfun_Id]));  | 
| 10751 | 1271  | 
qed "NSDERIV_Id";  | 
1272  | 
Addsimps [NSDERIV_Id];  | 
|
1273  | 
||
1274  | 
(*derivative of the identity function*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1275  | 
Goal "DERIV (%x. x) x :> 1";  | 
| 10751 | 1276  | 
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff RS sym]) 1);  | 
1277  | 
qed "DERIV_Id";  | 
|
1278  | 
Addsimps [DERIV_Id];  | 
|
1279  | 
||
1280  | 
bind_thm ("isCont_Id", DERIV_Id RS DERIV_isCont);
 | 
|
1281  | 
||
1282  | 
(*derivative of linear multiplication*)  | 
|
1283  | 
Goal "DERIV (op * c) x :> c";  | 
|
1284  | 
by (cut_inst_tac [("c","c"),("x","x")] (DERIV_Id RS DERIV_cmult) 1);
 | 
|
1285  | 
by (Asm_full_simp_tac 1);  | 
|
1286  | 
qed "DERIV_cmult_Id";  | 
|
1287  | 
Addsimps [DERIV_cmult_Id];  | 
|
1288  | 
||
1289  | 
Goal "NSDERIV (op * c) x :> c";  | 
|
1290  | 
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff]) 1);  | 
|
1291  | 
qed "NSDERIV_cmult_Id";  | 
|
1292  | 
Addsimps [NSDERIV_cmult_Id];  | 
|
1293  | 
||
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11468 
diff
changeset
 | 
1294  | 
Goal "DERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))";  | 
| 10751 | 1295  | 
by (induct_tac "n" 1);  | 
1296  | 
by (dtac (DERIV_Id RS DERIV_mult) 2);  | 
|
| 
10778
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
1297  | 
by (auto_tac (claset(),  | 
| 
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
1298  | 
simpset() addsimps [real_of_nat_Suc, real_add_mult_distrib]));  | 
| 10751 | 1299  | 
by (case_tac "0 < n" 1);  | 
1300  | 
by (dres_inst_tac [("x","x")] realpow_minus_mult 1);
 | 
|
| 
10778
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
1301  | 
by (auto_tac (claset(),  | 
| 
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
1302  | 
simpset() addsimps [real_mult_assoc, real_add_commute]));  | 
| 10751 | 1303  | 
qed "DERIV_pow";  | 
1304  | 
||
1305  | 
(* NS version *)  | 
|
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11468 
diff
changeset
 | 
1306  | 
Goal "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))";  | 
| 
10778
 
2c6605049646
more tidying, especially to remove real_of_posnat
 
paulson 
parents: 
10751 
diff
changeset
 | 
1307  | 
by (simp_tac (simpset() addsimps [NSDERIV_DERIV_iff, DERIV_pow]) 1);  | 
| 10751 | 1308  | 
qed "NSDERIV_pow";  | 
1309  | 
||
1310  | 
(*---------------------------------------------------------------  | 
|
1311  | 
Power of -1  | 
|
1312  | 
---------------------------------------------------------------*)  | 
|
1313  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1314  | 
(*Can't get rid of x \\<noteq> 0 because it isn't continuous at zero*)  | 
| 10751 | 1315  | 
Goalw [nsderiv_def]  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1316  | 
"x \\<noteq> 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))";  | 
| 10751 | 1317  | 
by (rtac ballI 1 THEN Asm_full_simp_tac 1 THEN Step_tac 1);  | 
| 12486 | 1318  | 
by (ftac Infinitesimal_add_not_zero 1);  | 
| 10751 | 1319  | 
by (asm_full_simp_tac (simpset() addsimps [hypreal_add_commute]) 2);  | 
1320  | 
by (auto_tac (claset(),  | 
|
1321  | 
simpset() addsimps [starfun_inverse_inverse, realpow_two]  | 
|
1322  | 
delsimps [hypreal_minus_mult_eq1 RS sym,  | 
|
1323  | 
hypreal_minus_mult_eq2 RS sym]));  | 
|
1324  | 
by (asm_full_simp_tac  | 
|
1325  | 
(simpset() addsimps [hypreal_inverse_add,  | 
|
1326  | 
hypreal_inverse_distrib RS sym, hypreal_minus_inverse RS sym]  | 
|
1327  | 
@ hypreal_add_ac @ hypreal_mult_ac  | 
|
1328  | 
delsimps [hypreal_minus_mult_eq1 RS sym,  | 
|
1329  | 
hypreal_minus_mult_eq2 RS sym] ) 1);  | 
|
1330  | 
by (asm_simp_tac (simpset() addsimps [hypreal_mult_assoc RS sym,  | 
|
1331  | 
hypreal_add_mult_distrib2]  | 
|
1332  | 
delsimps [hypreal_minus_mult_eq1 RS sym,  | 
|
1333  | 
hypreal_minus_mult_eq2 RS sym]) 1);  | 
|
1334  | 
by (res_inst_tac [("y"," inverse(- hypreal_of_real x * hypreal_of_real x)")] 
 | 
|
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1335  | 
approx_trans 1);  | 
| 
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1336  | 
by (rtac inverse_add_Infinitesimal_approx2 1);  | 
| 10751 | 1337  | 
by (auto_tac (claset() addSDs [hypreal_of_real_HFinite_diff_Infinitesimal],  | 
1338  | 
simpset() addsimps [hypreal_minus_inverse RS sym,  | 
|
1339  | 
HFinite_minus_iff]));  | 
|
1340  | 
by (rtac Infinitesimal_HFinite_mult2 1);  | 
|
1341  | 
by Auto_tac;  | 
|
1342  | 
qed "NSDERIV_inverse";  | 
|
1343  | 
||
1344  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1345  | 
Goal "x \\<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ Suc (Suc 0)))";  | 
| 10751 | 1346  | 
by (asm_simp_tac (simpset() addsimps [NSDERIV_inverse,  | 
1347  | 
NSDERIV_DERIV_iff RS sym] delsimps [realpow_Suc]) 1);  | 
|
1348  | 
qed "DERIV_inverse";  | 
|
1349  | 
||
1350  | 
(*--------------------------------------------------------------  | 
|
1351  | 
Derivative of inverse  | 
|
1352  | 
-------------------------------------------------------------*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1353  | 
Goal "[| DERIV f x :> d; f(x) \\<noteq> 0 |] \  | 
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11468 
diff
changeset
 | 
1354  | 
\ ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))";  | 
| 10751 | 1355  | 
by (rtac (real_mult_commute RS subst) 1);  | 
1356  | 
by (asm_simp_tac (simpset() addsimps [real_minus_mult_eq1,  | 
|
| 
12481
 
ea5d6da573c5
mods due to reorienting and renaming of real_minus_mult_eq1/2
 
nipkow 
parents: 
12018 
diff
changeset
 | 
1357  | 
realpow_inverse] delsimps [realpow_Suc, real_mult_minus_eq1]) 1);  | 
| 10751 | 1358  | 
by (fold_goals_tac [o_def]);  | 
1359  | 
by (blast_tac (claset() addSIs [DERIV_chain,DERIV_inverse]) 1);  | 
|
1360  | 
qed "DERIV_inverse_fun";  | 
|
1361  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1362  | 
Goal "[| NSDERIV f x :> d; f(x) \\<noteq> 0 |] \  | 
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11468 
diff
changeset
 | 
1363  | 
\ ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))";  | 
| 10751 | 1364  | 
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,  | 
1365  | 
DERIV_inverse_fun] delsimps [realpow_Suc]) 1);  | 
|
1366  | 
qed "NSDERIV_inverse_fun";  | 
|
1367  | 
||
1368  | 
(*--------------------------------------------------------------  | 
|
1369  | 
Derivative of quotient  | 
|
1370  | 
-------------------------------------------------------------*)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1371  | 
Goal "[| DERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> 0 |] \  | 
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11468 
diff
changeset
 | 
1372  | 
\ ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) + -(e*f(x))) / (g(x) ^ Suc (Suc 0))";  | 
| 10751 | 1373  | 
by (dres_inst_tac [("f","g")] DERIV_inverse_fun 1);
 | 
1374  | 
by (dtac DERIV_mult 2);  | 
|
1375  | 
by (REPEAT(assume_tac 1));  | 
|
1376  | 
by (asm_full_simp_tac  | 
|
1377  | 
(simpset() addsimps [real_divide_def, real_add_mult_distrib2,  | 
|
1378  | 
realpow_inverse,real_minus_mult_eq1] @ real_mult_ac  | 
|
| 
12481
 
ea5d6da573c5
mods due to reorienting and renaming of real_minus_mult_eq1/2
 
nipkow 
parents: 
12018 
diff
changeset
 | 
1379  | 
delsimps [realpow_Suc, real_mult_minus_eq1, real_mult_minus_eq2]) 1);  | 
| 10751 | 1380  | 
qed "DERIV_quotient";  | 
1381  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1382  | 
Goal "[| NSDERIV f x :> d; DERIV g x :> e; g(x) \\<noteq> 0 |] \  | 
| 10751 | 1383  | 
\ ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) \  | 
| 
11701
 
3d51fbf81c17
sane numerals (stage 1): added generic 1, removed 1' and 2 on nat,
 
wenzelm 
parents: 
11468 
diff
changeset
 | 
1384  | 
\ + -(e*f(x))) / (g(x) ^ Suc (Suc 0))";  | 
| 10751 | 1385  | 
by (asm_full_simp_tac (simpset() addsimps [NSDERIV_DERIV_iff,  | 
1386  | 
DERIV_quotient] delsimps [realpow_Suc]) 1);  | 
|
1387  | 
qed "NSDERIV_quotient";  | 
|
1388  | 
||
1389  | 
(* ------------------------------------------------------------------------ *)  | 
|
1390  | 
(* Caratheodory formulation of derivative at a point: standard proof *)  | 
|
1391  | 
(* ------------------------------------------------------------------------ *)  | 
|
1392  | 
||
1393  | 
Goal "(DERIV f x :> l) = \  | 
|
| 11383 | 1394  | 
\ (\\<exists>g. (\\<forall>z. f z - f x = g z * (z - x)) & isCont g x & g x = l)";  | 
| 11176 | 1395  | 
by Safe_tac;  | 
| 10751 | 1396  | 
by (res_inst_tac  | 
1397  | 
    [("x","%z. if  z = x then l else (f(z) - f(x)) / (z - x)")] exI 1);
 | 
|
1398  | 
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1399  | 
ARITH_PROVE "z \\<noteq> x ==> z - x \\<noteq> (0::real)"]));  | 
| 10751 | 1400  | 
by (auto_tac (claset(),simpset() addsimps [isCont_iff,DERIV_iff]));  | 
1401  | 
by (ALLGOALS(rtac (LIM_equal RS iffD1)));  | 
|
1402  | 
by (auto_tac (claset(),simpset() addsimps [real_diff_def,real_mult_assoc]));  | 
|
1403  | 
qed "CARAT_DERIV";  | 
|
1404  | 
||
1405  | 
Goal "NSDERIV f x :> l ==> \  | 
|
| 11383 | 1406  | 
\ \\<exists>g. (\\<forall>z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l";  | 
| 10751 | 1407  | 
by (auto_tac (claset(),simpset() addsimps [NSDERIV_DERIV_iff,  | 
1408  | 
isNSCont_isCont_iff,CARAT_DERIV]));  | 
|
1409  | 
qed "CARAT_NSDERIV";  | 
|
1410  | 
||
1411  | 
(* How about a NS proof? *)  | 
|
| 11383 | 1412  | 
Goal "(\\<forall>z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l \  | 
| 10751 | 1413  | 
\ ==> NSDERIV f x :> l";  | 
1414  | 
by (auto_tac (claset(),  | 
|
1415  | 
simpset() delsimprocs real_cancel_factor  | 
|
1416  | 
addsimps [NSDERIV_iff2]));  | 
|
1417  | 
by (auto_tac (claset(),  | 
|
1418  | 
simpset() addsimps [hypreal_mult_assoc]));  | 
|
1419  | 
by (asm_full_simp_tac (simpset() addsimps [hypreal_eq_minus_iff3 RS sym,  | 
|
1420  | 
hypreal_diff_def]) 1);  | 
|
1421  | 
by (asm_full_simp_tac (simpset() addsimps [isNSCont_def]) 1);  | 
|
1422  | 
qed "CARAT_DERIVD";  | 
|
1423  | 
||
1424  | 
||
1425  | 
||
1426  | 
(*--------------------------------------------------------------------------*)  | 
|
1427  | 
(* Lemmas about nested intervals and proof by bisection (cf.Harrison) *)  | 
|
1428  | 
(* All considerably tidied by lcp *)  | 
|
1429  | 
(*--------------------------------------------------------------------------*)  | 
|
1430  | 
||
| 11383 | 1431  | 
Goal "(\\<forall>n. (f::nat=>real) n \\<le> f (Suc n)) --> f m \\<le> f(m + no)";  | 
| 10751 | 1432  | 
by (induct_tac "no" 1);  | 
1433  | 
by (auto_tac (claset() addIs [order_trans], simpset()));  | 
|
1434  | 
qed_spec_mp "lemma_f_mono_add";  | 
|
1435  | 
||
| 11383 | 1436  | 
Goal "[| \\<forall>n. f(n) \\<le> f(Suc n); \  | 
1437  | 
\ \\<forall>n. g(Suc n) \\<le> g(n); \  | 
|
1438  | 
\ \\<forall>n. f(n) \\<le> g(n) |] \  | 
|
| 10751 | 1439  | 
\ ==> Bseq f";  | 
1440  | 
by (res_inst_tac [("k","f 0"),("K","g 0")] BseqI2 1 THEN rtac allI 1);
 | 
|
1441  | 
by (induct_tac "n" 1);  | 
|
1442  | 
by (auto_tac (claset() addIs [order_trans], simpset()));  | 
|
1443  | 
by (res_inst_tac [("y","g(Suc na)")] order_trans 1);
 | 
|
1444  | 
by (induct_tac "na" 2);  | 
|
1445  | 
by (auto_tac (claset() addIs [order_trans], simpset()));  | 
|
1446  | 
qed "f_inc_g_dec_Beq_f";  | 
|
1447  | 
||
| 11383 | 1448  | 
Goal "[| \\<forall>n. f(n) \\<le> f(Suc n); \  | 
1449  | 
\ \\<forall>n. g(Suc n) \\<le> g(n); \  | 
|
1450  | 
\ \\<forall>n. f(n) \\<le> g(n) |] \  | 
|
| 10751 | 1451  | 
\ ==> Bseq g";  | 
1452  | 
by (stac (Bseq_minus_iff RS sym) 1);  | 
|
1453  | 
by (res_inst_tac [("g","%x. -(f x)")] f_inc_g_dec_Beq_f 1); 
 | 
|
1454  | 
by Auto_tac;  | 
|
1455  | 
qed "f_inc_g_dec_Beq_g";  | 
|
1456  | 
||
| 11383 | 1457  | 
Goal "[| \\<forall>n. f n \\<le> f (Suc n); convergent f |] ==> f n \\<le> lim f";  | 
| 10751 | 1458  | 
by (rtac real_leI 1);  | 
1459  | 
by (auto_tac (claset(),  | 
|
1460  | 
simpset() addsimps [convergent_LIMSEQ_iff, LIMSEQ_iff, monoseq_Suc]));  | 
|
1461  | 
by (dtac real_less_sum_gt_zero 1);  | 
|
1462  | 
by (dres_inst_tac [("x","f n + - lim f")] spec 1);
 | 
|
1463  | 
by Safe_tac;  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1464  | 
by (dres_inst_tac [("P","%na. no\\<le>na --> ?Q na"),("x","no + n")] spec 1);
 | 
| 10751 | 1465  | 
by Auto_tac;  | 
| 11383 | 1466  | 
by (subgoal_tac "lim f \\<le> f(no + n)" 1);  | 
| 10751 | 1467  | 
by (induct_tac "no" 2);  | 
1468  | 
by (auto_tac (claset() addIs [order_trans],  | 
|
1469  | 
simpset() addsimps [real_diff_def, real_abs_def]));  | 
|
1470  | 
by (dres_inst_tac [("x","f(no + n)"),("no1","no")] 
 | 
|
1471  | 
(lemma_f_mono_add RSN (2,order_less_le_trans)) 1);  | 
|
1472  | 
by (auto_tac (claset(), simpset() addsimps [add_commute]));  | 
|
1473  | 
qed "f_inc_imp_le_lim";  | 
|
1474  | 
||
1475  | 
Goal "convergent g ==> lim (%x. - g x) = - (lim g)";  | 
|
1476  | 
by (rtac (LIMSEQ_minus RS limI) 1);  | 
|
1477  | 
by (asm_full_simp_tac (simpset() addsimps [convergent_LIMSEQ_iff]) 1);  | 
|
1478  | 
qed "lim_uminus";  | 
|
1479  | 
||
| 11383 | 1480  | 
Goal "[| \\<forall>n. g(Suc n) \\<le> g(n); convergent g |] ==> lim g \\<le> g n";  | 
1481  | 
by (subgoal_tac "- (g n) \\<le> - (lim g)" 1);  | 
|
| 10751 | 1482  | 
by (cut_inst_tac [("f", "%x. - (g x)")] f_inc_imp_le_lim 2);
 | 
1483  | 
by (auto_tac (claset(),  | 
|
1484  | 
simpset() addsimps [lim_uminus, convergent_minus_iff RS sym]));  | 
|
1485  | 
qed "g_dec_imp_lim_le";  | 
|
1486  | 
||
| 11383 | 1487  | 
Goal "[| \\<forall>n. f(n) \\<le> f(Suc n); \  | 
1488  | 
\ \\<forall>n. g(Suc n) \\<le> g(n); \  | 
|
1489  | 
\ \\<forall>n. f(n) \\<le> g(n) |] \  | 
|
1490  | 
\ ==> \\<exists>l m. l \\<le> m & ((\\<forall>n. f(n) \\<le> l) & f ----> l) & \  | 
|
1491  | 
\ ((\\<forall>n. m \\<le> g(n)) & g ----> m)";  | 
|
| 10751 | 1492  | 
by (subgoal_tac "monoseq f & monoseq g" 1);  | 
1493  | 
by (force_tac (claset(), simpset() addsimps [LIMSEQ_iff,monoseq_Suc]) 2);  | 
|
1494  | 
by (subgoal_tac "Bseq f & Bseq g" 1);  | 
|
1495  | 
by (blast_tac (claset() addIs [f_inc_g_dec_Beq_f, f_inc_g_dec_Beq_g]) 2);  | 
|
1496  | 
by (auto_tac (claset() addSDs [Bseq_monoseq_convergent],  | 
|
1497  | 
simpset() addsimps [convergent_LIMSEQ_iff]));  | 
|
1498  | 
by (res_inst_tac [("x","lim f")] exI 1);
 | 
|
1499  | 
by (res_inst_tac [("x","lim g")] exI 1);
 | 
|
1500  | 
by (auto_tac (claset() addIs [LIMSEQ_le], simpset()));  | 
|
1501  | 
by (auto_tac (claset(),  | 
|
1502  | 
simpset() addsimps [f_inc_imp_le_lim, g_dec_imp_lim_le,  | 
|
1503  | 
convergent_LIMSEQ_iff]));  | 
|
1504  | 
qed "lemma_nest";  | 
|
1505  | 
||
| 11383 | 1506  | 
Goal "[| \\<forall>n. f(n) \\<le> f(Suc n); \  | 
1507  | 
\ \\<forall>n. g(Suc n) \\<le> g(n); \  | 
|
1508  | 
\ \\<forall>n. f(n) \\<le> g(n); \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1509  | 
\ (%n. f(n) - g(n)) ----> 0 |] \  | 
| 11383 | 1510  | 
\ ==> \\<exists>l. ((\\<forall>n. f(n) \\<le> l) & f ----> l) & \  | 
1511  | 
\ ((\\<forall>n. l \\<le> g(n)) & g ----> l)";  | 
|
| 10751 | 1512  | 
by (dtac lemma_nest 1 THEN Auto_tac);  | 
1513  | 
by (subgoal_tac "l = m" 1);  | 
|
1514  | 
by (dres_inst_tac [("X","f")] LIMSEQ_diff 2);
 | 
|
1515  | 
by (auto_tac (claset() addIs [LIMSEQ_unique], simpset()));  | 
|
1516  | 
qed "lemma_nest_unique";  | 
|
1517  | 
||
1518  | 
||
| 11383 | 1519  | 
Goal "a \\<le> b ==> \  | 
1520  | 
\ \\<forall>n. fst (Bolzano_bisect P a b n) \\<le> snd (Bolzano_bisect P a b n)";  | 
|
| 10751 | 1521  | 
by (rtac allI 1);  | 
1522  | 
by (induct_tac "n" 1);  | 
|
1523  | 
by (auto_tac (claset(), simpset() addsimps [Let_def, split_def]));  | 
|
1524  | 
qed "Bolzano_bisect_le";  | 
|
1525  | 
||
| 11383 | 1526  | 
Goal "a \\<le> b ==> \  | 
1527  | 
\ \\<forall>n. fst(Bolzano_bisect P a b n) \\<le> fst (Bolzano_bisect P a b (Suc n))";  | 
|
| 10751 | 1528  | 
by (rtac allI 1);  | 
1529  | 
by (induct_tac "n" 1);  | 
|
1530  | 
by (auto_tac (claset(),  | 
|
1531  | 
simpset() addsimps [Bolzano_bisect_le, Let_def, split_def]));  | 
|
1532  | 
qed "Bolzano_bisect_fst_le_Suc";  | 
|
1533  | 
||
| 11383 | 1534  | 
Goal "a \\<le> b ==> \  | 
1535  | 
\ \\<forall>n. snd(Bolzano_bisect P a b (Suc n)) \\<le> snd (Bolzano_bisect P a b n)";  | 
|
| 10751 | 1536  | 
by (rtac allI 1);  | 
1537  | 
by (induct_tac "n" 1);  | 
|
1538  | 
by (auto_tac (claset(),  | 
|
1539  | 
simpset() addsimps [Bolzano_bisect_le, Let_def, split_def]));  | 
|
1540  | 
qed "Bolzano_bisect_Suc_le_snd";  | 
|
1541  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
1542  | 
Goal "((x::real) = y / (2 * z)) = (2 * x = y/z)";  | 
| 10751 | 1543  | 
by Auto_tac;  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1544  | 
by (dres_inst_tac [("f","%u. (1/2)*u")] arg_cong 1); 
 | 
| 10751 | 1545  | 
by Auto_tac;  | 
1546  | 
qed "eq_divide_2_times_iff";  | 
|
1547  | 
||
| 11383 | 1548  | 
Goal "a \\<le> b ==> \  | 
| 10751 | 1549  | 
\ snd(Bolzano_bisect P a b n) - fst(Bolzano_bisect P a b n) = \  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
1550  | 
\ (b-a) / (2 ^ n)";  | 
| 10751 | 1551  | 
by (induct_tac "n" 1);  | 
1552  | 
by (auto_tac (claset(),  | 
|
1553  | 
simpset() addsimps [eq_divide_2_times_iff, real_add_divide_distrib,  | 
|
1554  | 
Let_def, split_def]));  | 
|
1555  | 
by (auto_tac (claset(),  | 
|
1556  | 
simpset() addsimps (real_add_ac@[Bolzano_bisect_le, real_diff_def])));  | 
|
1557  | 
qed "Bolzano_bisect_diff";  | 
|
1558  | 
||
1559  | 
val Bolzano_nest_unique =  | 
|
1560  | 
[Bolzano_bisect_fst_le_Suc, Bolzano_bisect_Suc_le_snd, Bolzano_bisect_le]  | 
|
1561  | 
MRS lemma_nest_unique;  | 
|
1562  | 
||
1563  | 
(*P_prem is a looping simprule, so it works better if it isn't an assumption*)  | 
|
1564  | 
val P_prem::notP_prem::rest =  | 
|
| 11383 | 1565  | 
Goal "[| !!a b c. [| P(a,b); P(b,c); a \\<le> b; b \\<le> c|] ==> P(a,c); \  | 
1566  | 
\ ~ P(a,b); a \\<le> b |] ==> \  | 
|
| 10751 | 1567  | 
\ ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))";  | 
1568  | 
by (cut_facts_tac rest 1);  | 
|
1569  | 
by (induct_tac "n" 1);  | 
|
1570  | 
by (auto_tac (claset(),  | 
|
1571  | 
simpset() delsimps [surjective_pairing RS sym]  | 
|
1572  | 
addsimps [notP_prem, Let_def, split_def]));  | 
|
1573  | 
by (swap_res_tac [P_prem] 1);  | 
|
1574  | 
by (assume_tac 1);  | 
|
1575  | 
by (auto_tac (claset(), simpset() addsimps [Bolzano_bisect_le]));  | 
|
1576  | 
qed "not_P_Bolzano_bisect";  | 
|
1577  | 
||
1578  | 
(*Now we re-package P_prem as a formula*)  | 
|
| 11383 | 1579  | 
Goal "[| \\<forall>a b c. P(a,b) & P(b,c) & a \\<le> b & b \\<le> c --> P(a,c); \  | 
1580  | 
\ ~ P(a,b); a \\<le> b |] ==> \  | 
|
1581  | 
\ \\<forall>n. ~ P(fst(Bolzano_bisect P a b n), snd(Bolzano_bisect P a b n))";  | 
|
| 10751 | 1582  | 
by (blast_tac (claset() addSEs [not_P_Bolzano_bisect RSN (2,rev_notE)]) 1);  | 
1583  | 
qed "not_P_Bolzano_bisect'";  | 
|
1584  | 
||
1585  | 
||
| 11383 | 1586  | 
Goal "[| \\<forall>a b c. P(a,b) & P(b,c) & a \\<le> b & b \\<le> c --> P(a,c); \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1587  | 
\ \\<forall>x. \\<exists>d::real. 0 < d & \  | 
| 11383 | 1588  | 
\ (\\<forall>a b. a \\<le> x & x \\<le> b & (b - a) < d --> P(a,b)); \  | 
1589  | 
\ a \\<le> b |] \  | 
|
| 10751 | 1590  | 
\ ==> P(a,b)";  | 
1591  | 
by (rtac (inst "P1" "P" Bolzano_nest_unique RS exE) 1);  | 
|
1592  | 
by (REPEAT (assume_tac 1));  | 
|
1593  | 
by (rtac LIMSEQ_minus_cancel 1);  | 
|
1594  | 
by (asm_simp_tac (simpset() addsimps [Bolzano_bisect_diff,  | 
|
1595  | 
LIMSEQ_divide_realpow_zero]) 1);  | 
|
1596  | 
by (rtac ccontr 1);  | 
|
1597  | 
by (dtac not_P_Bolzano_bisect' 1);  | 
|
1598  | 
by (REPEAT (assume_tac 1));  | 
|
1599  | 
by (rename_tac "l" 1);  | 
|
1600  | 
by (dres_inst_tac [("x","l")] spec 1 THEN Clarify_tac 1);
 | 
|
1601  | 
by (rewtac LIMSEQ_def);  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1602  | 
by (dres_inst_tac [("P", "%r. 0<r --> ?Q r"), ("x","d/2")] spec 1);
 | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1603  | 
by (dres_inst_tac [("P", "%r. 0<r --> ?Q r"), ("x","d/2")] spec 1);
 | 
| 10751 | 1604  | 
by (dtac real_less_half_sum 1);  | 
1605  | 
by Safe_tac;  | 
|
1606  | 
(*linear arithmetic bug if we just use Asm_simp_tac*)  | 
|
1607  | 
by (ALLGOALS Asm_full_simp_tac);  | 
|
1608  | 
by (dres_inst_tac [("x","fst(Bolzano_bisect P a b (no + noa))")] spec 1);
 | 
|
1609  | 
by (dres_inst_tac [("x","snd(Bolzano_bisect P a b (no + noa))")] spec 1);
 | 
|
1610  | 
by Safe_tac;  | 
|
1611  | 
by (ALLGOALS Asm_simp_tac);  | 
|
1612  | 
by (res_inst_tac [("y","abs(fst(Bolzano_bisect P a b(no + noa)) - l) + \
 | 
|
1613  | 
\ abs(snd(Bolzano_bisect P a b(no + noa)) - l)")]  | 
|
1614  | 
order_le_less_trans 1);  | 
|
1615  | 
by (asm_simp_tac (simpset() addsimps [real_abs_def]) 1);  | 
|
1616  | 
by (rtac (real_sum_of_halves RS subst) 1);  | 
|
1617  | 
by (rtac real_add_less_mono 1);  | 
|
1618  | 
by (ALLGOALS  | 
|
1619  | 
(asm_full_simp_tac (simpset() addsimps [symmetric real_diff_def])));  | 
|
1620  | 
qed "lemma_BOLZANO";  | 
|
1621  | 
||
1622  | 
||
| 11383 | 1623  | 
Goal "((\\<forall>a b c. (a \\<le> b & b \\<le> c & P(a,b) & P(b,c)) --> P(a,c)) & \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1624  | 
\ (\\<forall>x. \\<exists>d::real. 0 < d & \  | 
| 11383 | 1625  | 
\ (\\<forall>a b. a \\<le> x & x \\<le> b & (b - a) < d --> P(a,b)))) \  | 
1626  | 
\ --> (\\<forall>a b. a \\<le> b --> P(a,b))";  | 
|
| 10751 | 1627  | 
by (Clarify_tac 1);  | 
1628  | 
by (blast_tac (claset() addIs [lemma_BOLZANO]) 1);  | 
|
1629  | 
qed "lemma_BOLZANO2";  | 
|
1630  | 
||
1631  | 
||
1632  | 
(*----------------------------------------------------------------------------*)  | 
|
1633  | 
(* Intermediate Value Theorem (prove contrapositive by bisection) *)  | 
|
1634  | 
(*----------------------------------------------------------------------------*)  | 
|
1635  | 
||
| 11383 | 1636  | 
Goal "[| f(a) \\<le> y & y \\<le> f(b); \  | 
1637  | 
\ a \\<le> b; \  | 
|
1638  | 
\ (\\<forall>x. a \\<le> x & x \\<le> b --> isCont f x) |] \  | 
|
1639  | 
\ ==> \\<exists>x. a \\<le> x & x \\<le> b & f(x) = y";  | 
|
| 10751 | 1640  | 
by (rtac contrapos_pp 1);  | 
1641  | 
by (assume_tac 1);  | 
|
1642  | 
by (cut_inst_tac  | 
|
| 11383 | 1643  | 
    [("P","%(u,v). a \\<le> u & u \\<le> v & v \\<le> b --> ~(f(u) \\<le> y & y \\<le> f(v))")] 
 | 
| 10751 | 1644  | 
lemma_BOLZANO2 1);  | 
| 11176 | 1645  | 
by Safe_tac;  | 
| 10751 | 1646  | 
by (ALLGOALS(Asm_full_simp_tac));  | 
1647  | 
by (Blast_tac 2);  | 
|
1648  | 
by (asm_full_simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);  | 
|
1649  | 
by (rtac ccontr 1);  | 
|
| 11383 | 1650  | 
by (subgoal_tac "a \\<le> x & x \\<le> b" 1);  | 
| 10751 | 1651  | 
by (Asm_full_simp_tac 2);  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1652  | 
by (dres_inst_tac [("P", "%d. 0<d --> ?P d"),("x","1")] spec 2);
 | 
| 10751 | 1653  | 
by (Step_tac 2);  | 
1654  | 
by (Asm_full_simp_tac 2);  | 
|
1655  | 
by (Asm_full_simp_tac 2);  | 
|
1656  | 
by (REPEAT(blast_tac (claset() addIs [order_trans]) 2));  | 
|
1657  | 
by (REPEAT(dres_inst_tac [("x","x")] spec 1));
 | 
|
1658  | 
by (Asm_full_simp_tac 1);  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1659  | 
by (dres_inst_tac [("P", "%r. ?P r --> (\\<exists>s. 0<s & ?Q r s)"),
 | 
| 10751 | 1660  | 
                   ("x","abs(y - f x)")] spec 1);
 | 
| 11176 | 1661  | 
by Safe_tac;  | 
| 10751 | 1662  | 
by (asm_full_simp_tac (simpset() addsimps []) 1);  | 
1663  | 
by (dres_inst_tac [("x","s")] spec 1);
 | 
|
1664  | 
by (Clarify_tac 1);  | 
|
1665  | 
by (cut_inst_tac [("R1.0","f x"),("R2.0","y")] real_linear 1);
 | 
|
1666  | 
by Safe_tac;  | 
|
1667  | 
by (dres_inst_tac [("x","ba - x")] spec 1);
 | 
|
1668  | 
by (ALLGOALS (asm_full_simp_tac (simpset() addsimps [abs_iff])));  | 
|
1669  | 
by (dres_inst_tac [("x","aa - x")] spec 1);
 | 
|
| 11383 | 1670  | 
by (case_tac "x \\<le> aa" 1);  | 
| 10751 | 1671  | 
by (ALLGOALS Asm_full_simp_tac);  | 
1672  | 
by (dres_inst_tac [("z","x"),("w","aa")] real_le_anti_sym 1);
 | 
|
1673  | 
by (assume_tac 1 THEN Asm_full_simp_tac 1);  | 
|
1674  | 
qed "IVT";  | 
|
1675  | 
||
1676  | 
||
| 11383 | 1677  | 
Goal "[| f(b) \\<le> y & y \\<le> f(a); \  | 
1678  | 
\ a \\<le> b; \  | 
|
1679  | 
\ (\\<forall>x. a \\<le> x & x \\<le> b --> isCont f x) \  | 
|
1680  | 
\ |] ==> \\<exists>x. a \\<le> x & x \\<le> b & f(x) = y";  | 
|
1681  | 
by (subgoal_tac "- f a \\<le> -y & -y \\<le> - f b" 1);  | 
|
1682  | 
by (thin_tac "f b \\<le> y & y \\<le> f a" 1);  | 
|
| 10751 | 1683  | 
by (dres_inst_tac [("f","%x. - f x")] IVT 1);
 | 
1684  | 
by (auto_tac (claset() addIs [isCont_minus],simpset()));  | 
|
1685  | 
qed "IVT2";  | 
|
1686  | 
||
1687  | 
||
1688  | 
(*HOL style here: object-level formulations*)  | 
|
| 11383 | 1689  | 
Goal "(f(a) \\<le> y & y \\<le> f(b) & a \\<le> b & \  | 
1690  | 
\ (\\<forall>x. a \\<le> x & x \\<le> b --> isCont f x)) \  | 
|
1691  | 
\ --> (\\<exists>x. a \\<le> x & x \\<le> b & f(x) = y)";  | 
|
| 10751 | 1692  | 
by (blast_tac (claset() addIs [IVT]) 1);  | 
1693  | 
qed "IVT_objl";  | 
|
1694  | 
||
| 11383 | 1695  | 
Goal "(f(b) \\<le> y & y \\<le> f(a) & a \\<le> b & \  | 
1696  | 
\ (\\<forall>x. a \\<le> x & x \\<le> b --> isCont f x)) \  | 
|
1697  | 
\ --> (\\<exists>x. a \\<le> x & x \\<le> b & f(x) = y)";  | 
|
| 10751 | 1698  | 
by (blast_tac (claset() addIs [IVT2]) 1);  | 
1699  | 
qed "IVT2_objl";  | 
|
1700  | 
||
1701  | 
(*---------------------------------------------------------------------------*)  | 
|
1702  | 
(* By bisection, function continuous on closed interval is bounded above *)  | 
|
1703  | 
(*---------------------------------------------------------------------------*)  | 
|
1704  | 
||
| 
10919
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1705  | 
Goal "abs (real x) = real (x::nat)";  | 
| 
 
144ede948e58
renamings: real_of_nat, real_of_int -> (overloaded) real
 
paulson 
parents: 
10834 
diff
changeset
 | 
1706  | 
by (auto_tac (claset() addIs [abs_eqI1], simpset()));  | 
| 10751 | 1707  | 
qed "abs_real_of_nat_cancel";  | 
1708  | 
Addsimps [abs_real_of_nat_cancel];  | 
|
1709  | 
||
| 
11713
 
883d559b0b8c
sane numerals (stage 3): provide generic "1" on all number types;
 
wenzelm 
parents: 
11704 
diff
changeset
 | 
1710  | 
Goal "~ abs(x) + (1::real) < x";  | 
| 10751 | 1711  | 
by (rtac real_leD 1);  | 
1712  | 
by (auto_tac (claset() addIs [abs_ge_self RS order_trans],simpset()));  | 
|
1713  | 
qed "abs_add_one_not_less_self";  | 
|
1714  | 
Addsimps [abs_add_one_not_less_self];  | 
|
1715  | 
||
1716  | 
||
| 11383 | 1717  | 
Goal "[| a \\<le> b; \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x |]\  | 
1718  | 
\ ==> \\<exists>M. \\<forall>x. a \\<le> x & x \\<le> b --> f(x) \\<le> M";  | 
|
1719  | 
by (cut_inst_tac [("P","%(u,v). a \\<le> u & u \\<le> v & v \\<le> b --> \
 | 
|
1720  | 
\ (\\<exists>M. \\<forall>x. u \\<le> x & x \\<le> v --> f x \\<le> M)")]  | 
|
| 10751 | 1721  | 
lemma_BOLZANO2 1);  | 
| 11176 | 1722  | 
by Safe_tac;  | 
1723  | 
by (ALLGOALS Asm_full_simp_tac);  | 
|
1724  | 
by (rename_tac "x xa ya M Ma" 1);  | 
|
1725  | 
by (cut_inst_tac [("x","M"),("y","Ma")] linorder_linear 1);
 | 
|
1726  | 
by Safe_tac;  | 
|
| 10751 | 1727  | 
by (res_inst_tac [("x","Ma")] exI 1);
 | 
| 11176 | 1728  | 
by (Clarify_tac 1);  | 
1729  | 
by (cut_inst_tac [("x","xb"),("y","xa")] linorder_linear 1);
 | 
|
1730  | 
by (Force_tac 1);  | 
|
| 10751 | 1731  | 
by (res_inst_tac [("x","M")] exI 1);
 | 
| 11176 | 1732  | 
by (Clarify_tac 1);  | 
1733  | 
by (cut_inst_tac [("x","xb"),("y","xa")] linorder_linear 1);
 | 
|
1734  | 
by (Force_tac 1);  | 
|
| 11383 | 1735  | 
by (case_tac "a \\<le> x & x \\<le> b" 1);  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1736  | 
by (res_inst_tac [("x","1")] exI 2);
 | 
| 11176 | 1737  | 
by (Force_tac 2);  | 
1738  | 
by (asm_full_simp_tac (simpset() addsimps [LIM_def,isCont_iff]) 1);  | 
|
| 10751 | 1739  | 
by (dres_inst_tac [("x","x")] spec 1 THEN Auto_tac);
 | 
| 11383 | 1740  | 
by (thin_tac "\\<forall>M. \\<exists>x. a \\<le> x & x \\<le> b & ~ f x \\<le> M" 1);  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1741  | 
by (dres_inst_tac [("x","1")] spec 1);
 | 
| 10751 | 1742  | 
by Auto_tac;  | 
| 11176 | 1743  | 
by (res_inst_tac [("x","s")] exI 1 THEN Clarify_tac 1);
 | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1744  | 
by (res_inst_tac [("x","abs(f x) + 1")] exI 1 THEN Clarify_tac 1);
 | 
| 11176 | 1745  | 
by (dres_inst_tac [("x","xa - x")] spec 1 THEN Safe_tac);
 | 
1746  | 
by (arith_tac 1);  | 
|
1747  | 
by (arith_tac 1);  | 
|
1748  | 
by (asm_full_simp_tac (simpset() addsimps [abs_ge_self]) 1);  | 
|
1749  | 
by (arith_tac 1);  | 
|
| 10751 | 1750  | 
qed "isCont_bounded";  | 
1751  | 
||
1752  | 
(*----------------------------------------------------------------------------*)  | 
|
1753  | 
(* Refine the above to existence of least upper bound *)  | 
|
1754  | 
(*----------------------------------------------------------------------------*)  | 
|
1755  | 
||
| 11383 | 1756  | 
Goal "((\\<exists>x. x \\<in> S) & (\\<exists>y. isUb UNIV S (y::real))) --> \  | 
1757  | 
\ (\\<exists>t. isLub UNIV S t)";  | 
|
| 10751 | 1758  | 
by (blast_tac (claset() addIs [reals_complete]) 1);  | 
1759  | 
qed "lemma_reals_complete";  | 
|
1760  | 
||
| 11383 | 1761  | 
Goal "[| a \\<le> b; \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x |] \  | 
1762  | 
\ ==> \\<exists>M. (\\<forall>x. a \\<le> x & x \\<le> b --> f(x) \\<le> M) & \  | 
|
1763  | 
\ (\\<forall>N. N < M --> (\\<exists>x. a \\<le> x & x \\<le> b & N < f(x)))";  | 
|
1764  | 
by (cut_inst_tac [("S","Collect (%y. \\<exists>x. a \\<le> x & x \\<le> b & y = f x)")]
 | 
|
| 10751 | 1765  | 
lemma_reals_complete 1);  | 
1766  | 
by Auto_tac;  | 
|
1767  | 
by (dtac isCont_bounded 1 THEN assume_tac 1);  | 
|
1768  | 
by (auto_tac (claset(),simpset() addsimps [isUb_def,leastP_def,  | 
|
1769  | 
isLub_def,setge_def,setle_def]));  | 
|
1770  | 
by (rtac exI 1 THEN Auto_tac);  | 
|
1771  | 
by (REPEAT(dtac spec 1) THEN Auto_tac);  | 
|
1772  | 
by (dres_inst_tac [("x","x")] spec 1);
 | 
|
1773  | 
by (auto_tac (claset() addSIs [real_leI],simpset()));  | 
|
1774  | 
qed "isCont_has_Ub";  | 
|
1775  | 
||
1776  | 
(*----------------------------------------------------------------------------*)  | 
|
1777  | 
(* Now show that it attains its upper bound *)  | 
|
1778  | 
(*----------------------------------------------------------------------------*)  | 
|
1779  | 
||
| 11383 | 1780  | 
Goal "[| a \\<le> b; \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x |] \  | 
1781  | 
\ ==> \\<exists>M. (\\<forall>x. a \\<le> x & x \\<le> b --> f(x) \\<le> M) & \  | 
|
1782  | 
\ (\\<exists>x. a \\<le> x & x \\<le> b & f(x) = M)";  | 
|
| 10751 | 1783  | 
by (ftac isCont_has_Ub 1 THEN assume_tac 1);  | 
1784  | 
by (Clarify_tac 1);  | 
|
1785  | 
by (res_inst_tac [("x","M")] exI 1);
 | 
|
1786  | 
by (Asm_full_simp_tac 1);  | 
|
1787  | 
by (rtac ccontr 1);  | 
|
| 11383 | 1788  | 
by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> f x < M" 1 THEN Step_tac 1);  | 
| 10751 | 1789  | 
by (rtac ccontr 2 THEN dtac real_leI 2);  | 
1790  | 
by (dres_inst_tac [("z","M")] real_le_anti_sym 2);
 | 
|
1791  | 
by (REPEAT(Blast_tac 2));  | 
|
| 11383 | 1792  | 
by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> isCont (%x. inverse(M - f x)) x" 1);  | 
| 11176 | 1793  | 
by Safe_tac;  | 
| 10751 | 1794  | 
by (EVERY[rtac isCont_inverse 2, rtac isCont_diff 2, rtac notI 4]);  | 
1795  | 
by (ALLGOALS(asm_full_simp_tac (simpset() addsimps [real_diff_eq_eq])));  | 
|
1796  | 
by (Blast_tac 2);  | 
|
1797  | 
by (subgoal_tac  | 
|
| 11383 | 1798  | 
"\\<exists>k. \\<forall>x. a \\<le> x & x \\<le> b --> (%x. inverse(M - (f x))) x \\<le> k" 1);  | 
| 10751 | 1799  | 
by (rtac isCont_bounded 2);  | 
| 11176 | 1800  | 
by Safe_tac;  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1801  | 
by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> 0 < inverse(M - f(x))" 1);  | 
| 10751 | 1802  | 
by (Asm_full_simp_tac 1);  | 
| 11176 | 1803  | 
by Safe_tac;  | 
| 10751 | 1804  | 
by (asm_full_simp_tac (simpset() addsimps [real_less_diff_eq]) 2);  | 
1805  | 
by (subgoal_tac  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1806  | 
"\\<forall>x. a \\<le> x & x \\<le> b --> (%x. inverse(M - (f x))) x < (k + 1)" 1);  | 
| 11176 | 1807  | 
by Safe_tac;  | 
| 10751 | 1808  | 
by (res_inst_tac [("y","k")] order_le_less_trans 2);
 | 
1809  | 
by (asm_full_simp_tac (simpset() addsimps [real_zero_less_one]) 3);  | 
|
1810  | 
by (Asm_full_simp_tac 2);  | 
|
| 11383 | 1811  | 
by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1812  | 
\ inverse(k + 1) < inverse((%x. inverse(M - (f x))) x)" 1);  | 
| 11176 | 1813  | 
by Safe_tac;  | 
| 10751 | 1814  | 
by (rtac real_inverse_less_swap 2);  | 
1815  | 
by (ALLGOALS Asm_full_simp_tac);  | 
|
1816  | 
by (dres_inst_tac [("P", "%N. N<M --> ?Q N"),
 | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1817  | 
                   ("x","M - inverse(k + 1)")] spec 1);
 | 
| 10751 | 1818  | 
by (Step_tac 1 THEN dtac real_leI 1);  | 
1819  | 
by (dtac (real_le_diff_eq RS iffD1) 1);  | 
|
1820  | 
by (REPEAT(dres_inst_tac [("x","a")] spec 1));
 | 
|
1821  | 
by (Asm_full_simp_tac 1);  | 
|
1822  | 
by (asm_full_simp_tac  | 
|
1823  | 
(simpset() addsimps [real_inverse_eq_divide, pos_real_divide_le_eq]) 1);  | 
|
1824  | 
by (cut_inst_tac [("x","k"),("y","M-f a")] real_0_less_mult_iff 1);
 | 
|
1825  | 
by (Asm_full_simp_tac 1);  | 
|
1826  | 
(*last one*)  | 
|
1827  | 
by (REPEAT(dres_inst_tac [("x","x")] spec 1));
 | 
|
1828  | 
by (Asm_full_simp_tac 1);  | 
|
1829  | 
qed "isCont_eq_Ub";  | 
|
1830  | 
||
1831  | 
||
1832  | 
(*----------------------------------------------------------------------------*)  | 
|
1833  | 
(* Same theorem for lower bound *)  | 
|
1834  | 
(*----------------------------------------------------------------------------*)  | 
|
1835  | 
||
| 11383 | 1836  | 
Goal "[| a \\<le> b; \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x |] \  | 
1837  | 
\ ==> \\<exists>M. (\\<forall>x. a \\<le> x & x \\<le> b --> M \\<le> f(x)) & \  | 
|
1838  | 
\ (\\<exists>x. a \\<le> x & x \\<le> b & f(x) = M)";  | 
|
1839  | 
by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> isCont (%x. -(f x)) x" 1);  | 
|
| 10751 | 1840  | 
by (blast_tac (claset() addIs [isCont_minus]) 2);  | 
1841  | 
by (dres_inst_tac [("f","(%x. -(f x))")] isCont_eq_Ub 1);
 | 
|
| 11176 | 1842  | 
by Safe_tac;  | 
| 10751 | 1843  | 
by Auto_tac;  | 
1844  | 
qed "isCont_eq_Lb";  | 
|
1845  | 
||
1846  | 
||
1847  | 
(* ------------------------------------------------------------------------- *)  | 
|
1848  | 
(* Another version. *)  | 
|
1849  | 
(* ------------------------------------------------------------------------- *)  | 
|
1850  | 
||
| 11383 | 1851  | 
Goal "[|a \\<le> b; \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x |] \  | 
1852  | 
\ ==> \\<exists>L M. (\\<forall>x. a \\<le> x & x \\<le> b --> L \\<le> f(x) & f(x) \\<le> M) & \  | 
|
1853  | 
\ (\\<forall>y. L \\<le> y & y \\<le> M --> (\\<exists>x. a \\<le> x & x \\<le> b & (f(x) = y)))";  | 
|
| 10751 | 1854  | 
by (ftac isCont_eq_Lb 1);  | 
1855  | 
by (ftac isCont_eq_Ub 2);  | 
|
1856  | 
by (REPEAT(assume_tac 1));  | 
|
| 11176 | 1857  | 
by Safe_tac;  | 
| 10751 | 1858  | 
by (res_inst_tac [("x","f x")] exI 1);
 | 
1859  | 
by (res_inst_tac [("x","f xa")] exI 1);
 | 
|
1860  | 
by (Asm_full_simp_tac 1);  | 
|
| 11176 | 1861  | 
by Safe_tac;  | 
1862  | 
by (cut_inst_tac [("x","x"),("y","xa")] linorder_linear 1);
 | 
|
1863  | 
by Safe_tac;  | 
|
| 10751 | 1864  | 
by (cut_inst_tac [("f","f"),("a","x"),("b","xa"),("y","y")] IVT_objl 1);
 | 
1865  | 
by (cut_inst_tac [("f","f"),("a","xa"),("b","x"),("y","y")] IVT2_objl 2);
 | 
|
| 11176 | 1866  | 
by Safe_tac;  | 
| 10751 | 1867  | 
by (res_inst_tac [("x","xb")] exI 2);
 | 
1868  | 
by (res_inst_tac [("x","xb")] exI 4);
 | 
|
1869  | 
by (ALLGOALS(Asm_full_simp_tac));  | 
|
1870  | 
qed "isCont_Lb_Ub";  | 
|
1871  | 
||
1872  | 
(*----------------------------------------------------------------------------*)  | 
|
1873  | 
(* If f'(x) > 0 then x is locally strictly increasing at the right *)  | 
|
1874  | 
(*----------------------------------------------------------------------------*)  | 
|
1875  | 
||
1876  | 
Goalw [deriv_def,LIM_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1877  | 
"[| DERIV f x :> l; 0 < l |] \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1878  | 
\ ==> \\<exists>d. 0 < d & (\\<forall>h. 0 < h & h < d --> f(x) < f(x + h))";  | 
| 10751 | 1879  | 
by (dtac spec 1 THEN Auto_tac);  | 
1880  | 
by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
 | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1881  | 
by (subgoal_tac "0 < l*h" 1);  | 
| 10751 | 1882  | 
by (asm_full_simp_tac (simpset() addsimps [real_0_less_mult_iff]) 2);  | 
1883  | 
by (dres_inst_tac [("x","h")] spec 1);
 | 
|
1884  | 
by (asm_full_simp_tac  | 
|
1885  | 
(simpset() addsimps [real_abs_def, real_inverse_eq_divide,  | 
|
1886  | 
pos_real_le_divide_eq, pos_real_less_divide_eq]  | 
|
1887  | 
addsplits [split_if_asm]) 1);  | 
|
1888  | 
qed "DERIV_left_inc";  | 
|
1889  | 
||
1890  | 
Goalw [deriv_def,LIM_def]  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1891  | 
"[| DERIV f x :> l; l < 0 |] ==> \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1892  | 
\ \\<exists>d. 0 < d & (\\<forall>h. 0 < h & h < d --> f(x) < f(x - h))";  | 
| 10751 | 1893  | 
by (dres_inst_tac [("x","-l")] spec 1 THEN Auto_tac);
 | 
1894  | 
by (res_inst_tac [("x","s")] exI 1 THEN Auto_tac);
 | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1895  | 
by (subgoal_tac "l*h < 0" 1);  | 
| 10751 | 1896  | 
by (asm_full_simp_tac (simpset() addsimps [real_mult_less_0_iff]) 2);  | 
1897  | 
by (dres_inst_tac [("x","-h")] spec 1);
 | 
|
1898  | 
by (asm_full_simp_tac  | 
|
1899  | 
(simpset() addsimps [real_abs_def, real_inverse_eq_divide,  | 
|
1900  | 
pos_real_less_divide_eq,  | 
|
1901  | 
symmetric real_diff_def]  | 
|
1902  | 
addsplits [split_if_asm]) 1);  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1903  | 
by (subgoal_tac "0 < (f (x - h) - f x)/h" 1);  | 
| 10751 | 1904  | 
by (arith_tac 2);  | 
1905  | 
by (asm_full_simp_tac  | 
|
1906  | 
(simpset() addsimps [pos_real_less_divide_eq]) 1);  | 
|
1907  | 
qed "DERIV_left_dec";  | 
|
1908  | 
||
1909  | 
||
1910  | 
Goal "[| DERIV f x :> l; \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1911  | 
\ \\<exists>d. 0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)) |] \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1912  | 
\ ==> l = 0";  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1913  | 
by (res_inst_tac [("R1.0","l"),("R2.0","0")] real_linear_less2 1);
 | 
| 11176 | 1914  | 
by Safe_tac;  | 
| 10751 | 1915  | 
by (dtac DERIV_left_dec 1);  | 
1916  | 
by (dtac DERIV_left_inc 3);  | 
|
| 11176 | 1917  | 
by Safe_tac;  | 
| 10751 | 1918  | 
by (dres_inst_tac [("d1.0","d"),("d2.0","da")] real_lbound_gt_zero 1);
 | 
1919  | 
by (dres_inst_tac [("d1.0","d"),("d2.0","da")] real_lbound_gt_zero 3);
 | 
|
| 11176 | 1920  | 
by Safe_tac;  | 
| 10751 | 1921  | 
by (dres_inst_tac [("x","x - e")] spec 1);
 | 
1922  | 
by (dres_inst_tac [("x","x + e")] spec 2);
 | 
|
1923  | 
by (auto_tac (claset(), simpset() addsimps [real_abs_def]));  | 
|
1924  | 
qed "DERIV_local_max";  | 
|
1925  | 
||
1926  | 
(*----------------------------------------------------------------------------*)  | 
|
1927  | 
(* Similar theorem for a local minimum *)  | 
|
1928  | 
(*----------------------------------------------------------------------------*)  | 
|
1929  | 
||
1930  | 
Goal "[| DERIV f x :> l; \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1931  | 
\ \\<exists>d::real. 0 < d & (\\<forall>y. abs(x - y) < d --> f(x) \\<le> f(y)) |] \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1932  | 
\ ==> l = 0";  | 
| 10751 | 1933  | 
by (dtac (DERIV_minus RS DERIV_local_max) 1);  | 
1934  | 
by Auto_tac;  | 
|
1935  | 
qed "DERIV_local_min";  | 
|
1936  | 
||
1937  | 
(*----------------------------------------------------------------------------*)  | 
|
1938  | 
(* In particular if a function is locally flat *)  | 
|
1939  | 
(*----------------------------------------------------------------------------*)  | 
|
1940  | 
||
1941  | 
Goal "[| DERIV f x :> l; \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1942  | 
\ \\<exists>d. 0 < d & (\\<forall>y. abs(x - y) < d --> f(x) = f(y)) |] \  | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1943  | 
\ ==> l = 0";  | 
| 10751 | 1944  | 
by (auto_tac (claset() addSDs [DERIV_local_max],simpset()));  | 
1945  | 
qed "DERIV_local_const";  | 
|
1946  | 
||
1947  | 
(*----------------------------------------------------------------------------*)  | 
|
1948  | 
(* Lemma about introducing open ball in open interval *)  | 
|
1949  | 
(*----------------------------------------------------------------------------*)  | 
|
1950  | 
||
1951  | 
Goal "[| a < x; x < b |] ==> \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1952  | 
\ \\<exists>d::real. 0 < d & (\\<forall>y. abs(x - y) < d --> a < y & y < b)";  | 
| 10751 | 1953  | 
by (simp_tac (simpset() addsimps [abs_interval_iff]) 1);  | 
| 11176 | 1954  | 
by (cut_inst_tac [("x","x - a"),("y","b - x")] linorder_linear 1);
 | 
1955  | 
by Safe_tac;  | 
|
| 10751 | 1956  | 
by (res_inst_tac [("x","x - a")] exI 1);
 | 
1957  | 
by (res_inst_tac [("x","b - x")] exI 2);
 | 
|
1958  | 
by Auto_tac;  | 
|
1959  | 
by (auto_tac (claset(),simpset() addsimps [real_less_diff_eq]));  | 
|
1960  | 
qed "lemma_interval_lt";  | 
|
1961  | 
||
1962  | 
Goal "[| a < x; x < b |] ==> \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1963  | 
\ \\<exists>d::real. 0 < d & (\\<forall>y. abs(x - y) < d --> a \\<le> y & y \\<le> b)";  | 
| 10751 | 1964  | 
by (dtac lemma_interval_lt 1);  | 
1965  | 
by Auto_tac;  | 
|
1966  | 
by (auto_tac (claset() addSIs [exI] ,simpset()));  | 
|
1967  | 
qed "lemma_interval";  | 
|
1968  | 
||
1969  | 
(*-----------------------------------------------------------------------  | 
|
1970  | 
Rolle's Theorem  | 
|
1971  | 
If f is defined and continuous on the finite closed interval [a,b]  | 
|
1972  | 
and differentiable a least on the open interval (a,b), and f(a) = f(b),  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1973  | 
then x0 \\<in> (a,b) such that f'(x0) = 0  | 
| 10751 | 1974  | 
----------------------------------------------------------------------*)  | 
1975  | 
||
1976  | 
Goal "[| a < b; f(a) = f(b); \  | 
|
| 11383 | 1977  | 
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \  | 
1978  | 
\ \\<forall>x. a < x & x < b --> f differentiable x \  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1979  | 
\ |] ==> \\<exists>z. a < z & z < b & DERIV f z :> 0";  | 
| 10751 | 1980  | 
by (ftac (order_less_imp_le RS isCont_eq_Ub) 1);  | 
1981  | 
by (EVERY1[assume_tac,Step_tac]);  | 
|
1982  | 
by (ftac (order_less_imp_le RS isCont_eq_Lb) 1);  | 
|
1983  | 
by (EVERY1[assume_tac,Step_tac]);  | 
|
1984  | 
by (case_tac "a < x & x < b" 1 THEN etac conjE 1);  | 
|
1985  | 
by (Asm_full_simp_tac 2);  | 
|
1986  | 
by (forw_inst_tac [("a","a"),("x","x")] lemma_interval 1);
 | 
|
1987  | 
by (EVERY1[assume_tac,etac exE]);  | 
|
1988  | 
by (res_inst_tac [("x","x")] exI 1 THEN Asm_full_simp_tac 1);
 | 
|
| 11383 | 1989  | 
by (subgoal_tac "(\\<exists>l. DERIV f x :> l) & \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
1990  | 
\ (\\<exists>d. 0 < d & (\\<forall>y. abs(x - y) < d --> f(y) \\<le> f(x)))" 1);  | 
| 10751 | 1991  | 
by (Clarify_tac 1 THEN rtac conjI 2);  | 
1992  | 
by (blast_tac (claset() addIs [differentiableD]) 2);  | 
|
1993  | 
by (Blast_tac 2);  | 
|
1994  | 
by (ftac DERIV_local_max 1);  | 
|
1995  | 
by (EVERY1[Blast_tac,Blast_tac]);  | 
|
1996  | 
by (case_tac "a < xa & xa < b" 1 THEN etac conjE 1);  | 
|
1997  | 
by (Asm_full_simp_tac 2);  | 
|
1998  | 
by (forw_inst_tac [("a","a"),("x","xa")] lemma_interval 1);
 | 
|
1999  | 
by (EVERY1[assume_tac,etac exE]);  | 
|
2000  | 
by (res_inst_tac [("x","xa")] exI 1 THEN Asm_full_simp_tac 1);
 | 
|
| 11383 | 2001  | 
by (subgoal_tac "(\\<exists>l. DERIV f xa :> l) & \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2002  | 
\ (\\<exists>d. 0 < d & (\\<forall>y. abs(xa - y) < d --> f(xa) \\<le> f(y)))" 1);  | 
| 10751 | 2003  | 
by (Clarify_tac 1 THEN rtac conjI 2);  | 
2004  | 
by (blast_tac (claset() addIs [differentiableD]) 2);  | 
|
2005  | 
by (Blast_tac 2);  | 
|
2006  | 
by (ftac DERIV_local_min 1);  | 
|
2007  | 
by (EVERY1[Blast_tac,Blast_tac]);  | 
|
| 11383 | 2008  | 
by (subgoal_tac "\\<forall>x. a \\<le> x & x \\<le> b --> f(x) = f(b)" 1);  | 
| 10751 | 2009  | 
by (Clarify_tac 2);  | 
2010  | 
by (rtac real_le_anti_sym 2);  | 
|
2011  | 
by (subgoal_tac "f b = f x" 2);  | 
|
2012  | 
by (Asm_full_simp_tac 2);  | 
|
2013  | 
by (res_inst_tac [("x1","a"),("y1","x")] (order_le_imp_less_or_eq RS disjE) 2);
 | 
|
2014  | 
by (assume_tac 2);  | 
|
2015  | 
by (dres_inst_tac [("z","x"),("w","b")] real_le_anti_sym 2);
 | 
|
2016  | 
by (subgoal_tac "f b = f xa" 5);  | 
|
2017  | 
by (Asm_full_simp_tac 5);  | 
|
2018  | 
by (res_inst_tac [("x1","a"),("y1","xa")] (order_le_imp_less_or_eq RS disjE) 5);
 | 
|
2019  | 
by (assume_tac 5);  | 
|
2020  | 
by (dres_inst_tac [("z","xa"),("w","b")] real_le_anti_sym 5);
 | 
|
2021  | 
by (REPEAT(Asm_full_simp_tac 2));  | 
|
2022  | 
by (dtac real_dense 1 THEN etac exE 1);  | 
|
2023  | 
by (res_inst_tac [("x","r")] exI 1 THEN Asm_full_simp_tac 1);
 | 
|
2024  | 
by (etac conjE 1);  | 
|
2025  | 
by (forw_inst_tac [("a","a"),("x","r")] lemma_interval 1);
 | 
|
2026  | 
by (EVERY1[assume_tac, etac exE]);  | 
|
| 11383 | 2027  | 
by (subgoal_tac "(\\<exists>l. DERIV f r :> l) & \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2028  | 
\ (\\<exists>d. 0 < d & (\\<forall>y. abs(r - y) < d --> f(r) = f(y)))" 1);  | 
| 10751 | 2029  | 
by (Clarify_tac 1 THEN rtac conjI 2);  | 
2030  | 
by (blast_tac (claset() addIs [differentiableD]) 2);  | 
|
2031  | 
by (EVERY1[ftac DERIV_local_const, Blast_tac, Blast_tac]);  | 
|
2032  | 
by (res_inst_tac [("x","d")] exI 1);
 | 
|
2033  | 
by (EVERY1[rtac conjI, Blast_tac, rtac allI, rtac impI]);  | 
|
2034  | 
by (res_inst_tac [("s","f b")] trans 1);
 | 
|
2035  | 
by (blast_tac (claset() addSDs [order_less_imp_le]) 1);  | 
|
2036  | 
by (rtac sym 1 THEN Blast_tac 1);  | 
|
2037  | 
qed "Rolle";  | 
|
2038  | 
||
2039  | 
(*----------------------------------------------------------------------------*)  | 
|
2040  | 
(* Mean value theorem *)  | 
|
2041  | 
(*----------------------------------------------------------------------------*)  | 
|
2042  | 
||
2043  | 
Goal "f a - (f b - f a)/(b - a) * a = \  | 
|
2044  | 
\ f b - (f b - f a)/(b - a) * (b::real)";  | 
|
2045  | 
by (case_tac "a = b" 1);  | 
|
2046  | 
by (asm_full_simp_tac (simpset() addsimps [REAL_DIVIDE_ZERO]) 1);  | 
|
2047  | 
by (res_inst_tac [("c1","b - a")] (real_mult_left_cancel RS iffD1) 1);
 | 
|
2048  | 
by (arith_tac 1);  | 
|
2049  | 
by (auto_tac (claset(),  | 
|
2050  | 
simpset() addsimps [real_diff_mult_distrib2]));  | 
|
2051  | 
by (auto_tac (claset(),  | 
|
2052  | 
simpset() addsimps [real_diff_mult_distrib]));  | 
|
2053  | 
qed "lemma_MVT";  | 
|
2054  | 
||
2055  | 
Goal "[| a < b; \  | 
|
| 11383 | 2056  | 
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \  | 
2057  | 
\ \\<forall>x. a < x & x < b --> f differentiable x |] \  | 
|
2058  | 
\ ==> \\<exists>l z. a < z & z < b & DERIV f z :> l & \  | 
|
| 10751 | 2059  | 
\ (f(b) - f(a) = (b - a) * l)";  | 
2060  | 
by (dres_inst_tac [("f","%x. f(x) - (((f(b) - f(a)) / (b - a)) * x)")]
 | 
|
2061  | 
Rolle 1);  | 
|
2062  | 
by (rtac lemma_MVT 1);  | 
|
| 11176 | 2063  | 
by Safe_tac;  | 
| 10751 | 2064  | 
by (rtac isCont_diff 1 THEN Blast_tac 1);  | 
2065  | 
by (rtac (isCont_const RS isCont_mult) 1);  | 
|
2066  | 
by (rtac isCont_Id 1);  | 
|
2067  | 
by (dres_inst_tac [("P", "%x. ?Pre x --> f differentiable x"), 
 | 
|
2068  | 
                   ("x","x")] spec 1);
 | 
|
2069  | 
by (asm_full_simp_tac (simpset() addsimps [differentiable_def]) 1);  | 
|
| 11176 | 2070  | 
by Safe_tac;  | 
| 10751 | 2071  | 
by (res_inst_tac [("x","xa - ((f(b) - f(a)) / (b - a))")] exI 1);
 | 
2072  | 
by (rtac DERIV_diff 1 THEN assume_tac 1);  | 
|
2073  | 
(*derivative of a linear function is the constant...*)  | 
|
2074  | 
by (subgoal_tac "(%x. (f b - f a) * x / (b - a)) = \  | 
|
2075  | 
\ op * ((f b - f a) / (b - a))" 1);  | 
|
2076  | 
by (rtac ext 2 THEN Simp_tac 2);  | 
|
2077  | 
by (Asm_full_simp_tac 1);  | 
|
2078  | 
(*final case*)  | 
|
2079  | 
by (res_inst_tac [("x","((f(b) - f(a)) / (b - a))")] exI 1);
 | 
|
2080  | 
by (res_inst_tac [("x","z")] exI 1);
 | 
|
| 11176 | 2081  | 
by Safe_tac;  | 
| 10751 | 2082  | 
by (Asm_full_simp_tac 2);  | 
2083  | 
by (subgoal_tac "DERIV (%x. ((f(b) - f(a)) / (b - a)) * x) z :> \  | 
|
2084  | 
\ ((f(b) - f(a)) / (b - a))" 1);  | 
|
2085  | 
by (rtac DERIV_cmult_Id 2);  | 
|
2086  | 
by (dtac DERIV_add 1 THEN assume_tac 1);  | 
|
2087  | 
by (asm_full_simp_tac (simpset() addsimps [real_add_assoc, real_diff_def]) 1);  | 
|
2088  | 
qed "MVT";  | 
|
2089  | 
||
2090  | 
(*----------------------------------------------------------------------------*)  | 
|
2091  | 
(* Theorem that function is constant if its derivative is 0 over an interval. *)  | 
|
2092  | 
(*----------------------------------------------------------------------------*)  | 
|
2093  | 
||
2094  | 
Goal "[| a < b; \  | 
|
| 11383 | 2095  | 
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2096  | 
\ \\<forall>x. a < x & x < b --> DERIV f x :> 0 |] \  | 
| 10751 | 2097  | 
\ ==> (f b = f a)";  | 
2098  | 
by (dtac MVT 1 THEN assume_tac 1);  | 
|
2099  | 
by (blast_tac (claset() addIs [differentiableI]) 1);  | 
|
2100  | 
by (auto_tac (claset() addSDs [DERIV_unique],simpset()  | 
|
2101  | 
addsimps [real_diff_eq_eq]));  | 
|
2102  | 
qed "DERIV_isconst_end";  | 
|
2103  | 
||
2104  | 
Goal "[| a < b; \  | 
|
| 11383 | 2105  | 
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2106  | 
\ \\<forall>x. a < x & x < b --> DERIV f x :> 0 |] \  | 
| 11383 | 2107  | 
\ ==> \\<forall>x. a \\<le> x & x \\<le> b --> f x = f a";  | 
| 11176 | 2108  | 
by Safe_tac;  | 
| 10751 | 2109  | 
by (dres_inst_tac [("x","a")] order_le_imp_less_or_eq 1);
 | 
| 11176 | 2110  | 
by Safe_tac;  | 
| 10751 | 2111  | 
by (dres_inst_tac [("b","x")] DERIV_isconst_end 1);
 | 
2112  | 
by Auto_tac;  | 
|
2113  | 
qed "DERIV_isconst1";  | 
|
2114  | 
||
2115  | 
Goal "[| a < b; \  | 
|
| 11383 | 2116  | 
\ \\<forall>x. a \\<le> x & x \\<le> b --> isCont f x; \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2117  | 
\ \\<forall>x. a < x & x < b --> DERIV f x :> 0; \  | 
| 11383 | 2118  | 
\ a \\<le> x; x \\<le> b |] \  | 
| 10751 | 2119  | 
\ ==> f x = f a";  | 
2120  | 
by (blast_tac (claset() addDs [DERIV_isconst1]) 1);  | 
|
2121  | 
qed "DERIV_isconst2";  | 
|
2122  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2123  | 
Goal "\\<forall>x. DERIV f x :> 0 ==> f(x) = f(y)";  | 
| 10751 | 2124  | 
by (res_inst_tac [("R1.0","x"),("R2.0","y")] real_linear_less2 1);
 | 
2125  | 
by (rtac sym 1);  | 
|
2126  | 
by (auto_tac (claset() addIs [DERIV_isCont,DERIV_isconst_end],simpset()));  | 
|
2127  | 
qed "DERIV_isconst_all";  | 
|
2128  | 
||
| 11383 | 2129  | 
Goal "[|a \\<noteq> b; \\<forall>x. DERIV f x :> k |] ==> (f(b) - f(a)) = (b - a) * k";  | 
| 10751 | 2130  | 
by (res_inst_tac [("R1.0","a"),("R2.0","b")] real_linear_less2 1);
 | 
2131  | 
by Auto_tac;  | 
|
2132  | 
by (ALLGOALS(dres_inst_tac [("f","f")] MVT));
 | 
|
2133  | 
by (auto_tac (claset() addDs [DERIV_isCont,DERIV_unique],simpset() addsimps  | 
|
2134  | 
[differentiable_def]));  | 
|
2135  | 
by (auto_tac (claset() addDs [DERIV_unique],  | 
|
| 
12481
 
ea5d6da573c5
mods due to reorienting and renaming of real_minus_mult_eq1/2
 
nipkow 
parents: 
12018 
diff
changeset
 | 
2136  | 
simpset() addsimps [real_add_mult_distrib, real_diff_def]));  | 
| 10751 | 2137  | 
qed "DERIV_const_ratio_const";  | 
2138  | 
||
| 11383 | 2139  | 
Goal "[|a \\<noteq> b; \\<forall>x. DERIV f x :> k |] ==> (f(b) - f(a))/(b - a) = k";  | 
| 10751 | 2140  | 
by (res_inst_tac [("c1","b - a")] (real_mult_right_cancel RS iffD1) 1);
 | 
2141  | 
by (auto_tac (claset() addSDs [DERIV_const_ratio_const],  | 
|
2142  | 
simpset() addsimps [real_mult_assoc]));  | 
|
2143  | 
qed "DERIV_const_ratio_const2";  | 
|
2144  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
2145  | 
Goal "((a + b) /2 - a) = (b - a)/(2::real)";  | 
| 10751 | 2146  | 
by Auto_tac;  | 
2147  | 
qed "real_average_minus_first";  | 
|
2148  | 
Addsimps [real_average_minus_first];  | 
|
2149  | 
||
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
2150  | 
Goal "((b + a)/2 - a) = (b - a)/(2::real)";  | 
| 10751 | 2151  | 
by Auto_tac;  | 
2152  | 
qed "real_average_minus_second";  | 
|
2153  | 
Addsimps [real_average_minus_second];  | 
|
2154  | 
||
2155  | 
||
2156  | 
(* Gallileo's "trick": average velocity = av. of end velocities *)  | 
|
| 11383 | 2157  | 
Goal "[|a \\<noteq> (b::real); \\<forall>x. DERIV v x :> k|] \  | 
| 
11704
 
3c50a2cd6f00
* sane numerals (stage 2): plain "num" syntax (removed "#");
 
wenzelm 
parents: 
11701 
diff
changeset
 | 
2158  | 
\ ==> v((a + b)/2) = (v a + v b)/2";  | 
| 10751 | 2159  | 
by (res_inst_tac [("R1.0","a"),("R2.0","b")] real_linear_less2 1);
 | 
2160  | 
by Auto_tac;  | 
|
2161  | 
by (ftac DERIV_const_ratio_const2 1 THEN assume_tac 1);  | 
|
2162  | 
by (ftac DERIV_const_ratio_const2 2 THEN assume_tac 2);  | 
|
2163  | 
by (dtac real_less_half_sum 1);  | 
|
2164  | 
by (dtac real_gt_half_sum 2);  | 
|
2165  | 
by (ftac (real_not_refl2 RS DERIV_const_ratio_const2) 1 THEN assume_tac 1);  | 
|
2166  | 
by (dtac ((real_not_refl2 RS not_sym) RS DERIV_const_ratio_const2) 2  | 
|
2167  | 
THEN assume_tac 2);  | 
|
2168  | 
by (ALLGOALS (dres_inst_tac [("f","%u. (b-a)*u")] arg_cong)); 
 | 
|
2169  | 
by (auto_tac (claset(), simpset() addsimps [real_inverse_eq_divide]));  | 
|
2170  | 
by (asm_full_simp_tac (simpset() addsimps [real_add_commute, eq_commute]) 1);  | 
|
2171  | 
qed "DERIV_const_average";  | 
|
2172  | 
||
2173  | 
||
2174  | 
(* ------------------------------------------------------------------------ *)  | 
|
2175  | 
(* Dull lemma that an continuous injection on an interval must have a strict*)  | 
|
2176  | 
(* maximum at an end point, not in the middle. *)  | 
|
2177  | 
(* ------------------------------------------------------------------------ *)  | 
|
2178  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2179  | 
Goal "[|0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \  | 
| 11383 | 2180  | 
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \  | 
2181  | 
\ ==> ~(\\<forall>z. abs(z - x) \\<le> d --> f(z) \\<le> f(x))";  | 
|
| 10751 | 2182  | 
by (rtac notI 1);  | 
2183  | 
by (rotate_tac 3 1);  | 
|
2184  | 
by (forw_inst_tac [("x","x - d")] spec 1);
 | 
|
2185  | 
by (forw_inst_tac [("x","x + d")] spec 1);
 | 
|
| 11176 | 2186  | 
by Safe_tac;  | 
| 10751 | 2187  | 
by (cut_inst_tac [("x","f(x - d)"),("y","f(x + d)")] 
 | 
| 11383 | 2188  | 
(ARITH_PROVE "x \\<le> y | y \\<le> (x::real)") 4);  | 
| 10751 | 2189  | 
by (etac disjE 4);  | 
2190  | 
by (REPEAT(arith_tac 1));  | 
|
2191  | 
by (cut_inst_tac [("f","f"),("a","x - d"),("b","x"),("y","f(x + d)")]
 | 
|
2192  | 
IVT_objl 1);  | 
|
| 11176 | 2193  | 
by Safe_tac;  | 
| 10751 | 2194  | 
by (arith_tac 1);  | 
2195  | 
by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);  | 
|
2196  | 
by (dres_inst_tac [("f","g")] arg_cong 1);
 | 
|
2197  | 
by (rotate_tac 2 1);  | 
|
2198  | 
by (forw_inst_tac [("x","xa")] spec 1);
 | 
|
2199  | 
by (dres_inst_tac [("x","x + d")] spec 1);
 | 
|
2200  | 
by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);  | 
|
2201  | 
(* 2nd case: similar *)  | 
|
2202  | 
by (cut_inst_tac [("f","f"),("a","x"),("b","x + d"),("y","f(x - d)")]
 | 
|
2203  | 
IVT2_objl 1);  | 
|
| 11176 | 2204  | 
by Safe_tac;  | 
| 10751 | 2205  | 
by (arith_tac 1);  | 
2206  | 
by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);  | 
|
2207  | 
by (dres_inst_tac [("f","g")] arg_cong 1);
 | 
|
2208  | 
by (rotate_tac 2 1);  | 
|
2209  | 
by (forw_inst_tac [("x","xa")] spec 1);
 | 
|
2210  | 
by (dres_inst_tac [("x","x - d")] spec 1);
 | 
|
2211  | 
by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);  | 
|
2212  | 
qed "lemma_isCont_inj";  | 
|
2213  | 
||
2214  | 
(* ------------------------------------------------------------------------ *)  | 
|
2215  | 
(* Similar version for lower bound *)  | 
|
2216  | 
(* ------------------------------------------------------------------------ *)  | 
|
2217  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2218  | 
Goal "[|0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \  | 
| 11383 | 2219  | 
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \  | 
2220  | 
\ ==> ~(\\<forall>z. abs(z - x) \\<le> d --> f(x) \\<le> f(z))";  | 
|
| 10751 | 2221  | 
by (auto_tac (claset() addSDs [(asm_full_simplify (simpset())  | 
2222  | 
    (read_instantiate [("f","%x. - f x"),("g","%y. g(-y)"),("x","x"),("d","d")]
 | 
|
2223  | 
lemma_isCont_inj))],simpset() addsimps [isCont_minus]));  | 
|
2224  | 
qed "lemma_isCont_inj2";  | 
|
2225  | 
||
2226  | 
(* ------------------------------------------------------------------------ *)  | 
|
2227  | 
(* Show there's an interval surrounding f(x) in f[[x - d, x + d]] *)  | 
|
2228  | 
(* Also from John's theory *)  | 
|
2229  | 
(* ------------------------------------------------------------------------ *)  | 
|
2230  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2231  | 
val lemma_le = ARITH_PROVE "0 \\<le> (d::real) ==> -d \\<le> d";  | 
| 10751 | 2232  | 
|
2233  | 
(* FIXME: awful proof - needs improvement *)  | 
|
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2234  | 
Goal "[| 0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f z) = z; \  | 
| 11383 | 2235  | 
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2236  | 
\ ==> \\<exists>e. 0 < e & \  | 
| 11383 | 2237  | 
\ (\\<forall>y. \  | 
2238  | 
\ abs(y - f(x)) \\<le> e --> \  | 
|
2239  | 
\ (\\<exists>z. abs(z - x) \\<le> d & (f z = y)))";  | 
|
| 10751 | 2240  | 
by (ftac order_less_imp_le 1);  | 
2241  | 
by (dtac (lemma_le RS (asm_full_simplify (simpset()) (read_instantiate  | 
|
2242  | 
    [("f","f"),("a","x - d"),("b","x + d")] isCont_Lb_Ub))) 1);
 | 
|
| 11176 | 2243  | 
by Safe_tac;  | 
| 10751 | 2244  | 
by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);  | 
| 11383 | 2245  | 
by (subgoal_tac "L \\<le> f x & f x \\<le> M" 1);  | 
| 10751 | 2246  | 
by (dres_inst_tac [("P", "%v. ?P v --> ?Q v & ?R v"), ("x","x")] spec 2);
 | 
2247  | 
by (Asm_full_simp_tac 2);  | 
|
2248  | 
by (subgoal_tac "L < f x & f x < M" 1);  | 
|
| 11176 | 2249  | 
by Safe_tac;  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2250  | 
by (dres_inst_tac [("x","L")] (ARITH_PROVE "x < y ==> 0 < y - (x::real)") 1);
 | 
| 
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2251  | 
by (dres_inst_tac [("x","f x")] (ARITH_PROVE "x < y ==> 0 < y - (x::real)") 1);
 | 
| 10751 | 2252  | 
by (dres_inst_tac [("d1.0","f x - L"),("d2.0","M - f x")] 
 | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2253  | 
(real_lbound_gt_zero) 1);  | 
| 11176 | 2254  | 
by Safe_tac;  | 
| 10751 | 2255  | 
by (res_inst_tac [("x","e")] exI 1);
 | 
| 11176 | 2256  | 
by Safe_tac;  | 
| 10751 | 2257  | 
by (asm_full_simp_tac (simpset() addsimps [abs_le_interval_iff]) 1);  | 
| 11383 | 2258  | 
by (dres_inst_tac [("P","%v. ?PP v --> (\\<exists>xa. ?Q v xa)"),("x","y")] spec 1);
 | 
| 10751 | 2259  | 
by (Step_tac 1 THEN REPEAT(arith_tac 1));  | 
2260  | 
by (res_inst_tac [("x","xa")] exI 1);
 | 
|
2261  | 
by (arith_tac 1);  | 
|
| 11383 | 2262  | 
by (ALLGOALS(etac (ARITH_PROVE "[|x \\<le> y; x \\<noteq> y |] ==> x < (y::real)")));  | 
| 10751 | 2263  | 
by (ALLGOALS(rotate_tac 3));  | 
2264  | 
by (dtac lemma_isCont_inj2 1);  | 
|
2265  | 
by (assume_tac 2);  | 
|
2266  | 
by (dtac lemma_isCont_inj 3);  | 
|
2267  | 
by (assume_tac 4);  | 
|
2268  | 
by (TRYALL(assume_tac));  | 
|
| 11176 | 2269  | 
by Safe_tac;  | 
| 10751 | 2270  | 
by (ALLGOALS(dres_inst_tac [("x","z")] spec));
 | 
2271  | 
by (ALLGOALS(arith_tac));  | 
|
2272  | 
qed "isCont_inj_range";  | 
|
2273  | 
||
2274  | 
||
2275  | 
(* ------------------------------------------------------------------------ *)  | 
|
2276  | 
(* Continuity of inverse function *)  | 
|
2277  | 
(* ------------------------------------------------------------------------ *)  | 
|
2278  | 
||
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2279  | 
Goal "[| 0 < d; \\<forall>z. abs(z - x) \\<le> d --> g(f(z)) = z; \  | 
| 11383 | 2280  | 
\ \\<forall>z. abs(z - x) \\<le> d --> isCont f z |] \  | 
| 10751 | 2281  | 
\ ==> isCont g (f x)";  | 
2282  | 
by (simp_tac (simpset() addsimps [isCont_iff,LIM_def]) 1);  | 
|
| 11176 | 2283  | 
by Safe_tac;  | 
| 
12018
 
ec054019c910
Numerals and simprocs for types real and hypreal.  The abstract
 
paulson 
parents: 
11713 
diff
changeset
 | 
2284  | 
by (dres_inst_tac [("d1.0","r")] (real_lbound_gt_zero) 1);
 | 
| 10751 | 2285  | 
by (assume_tac 1 THEN Step_tac 1);  | 
| 11383 | 2286  | 
by (subgoal_tac "\\<forall>z. abs(z - x) \\<le> e --> (g(f z) = z)" 1);  | 
| 10751 | 2287  | 
by (Force_tac 2);  | 
| 11383 | 2288  | 
by (subgoal_tac "\\<forall>z. abs(z - x) \\<le> e --> isCont f z" 1);  | 
| 10751 | 2289  | 
by (Force_tac 2);  | 
2290  | 
by (dres_inst_tac [("d","e")] isCont_inj_range 1);
 | 
|
2291  | 
by (assume_tac 2 THEN assume_tac 1);  | 
|
| 11176 | 2292  | 
by Safe_tac;  | 
| 10751 | 2293  | 
by (res_inst_tac [("x","ea")] exI 1);
 | 
2294  | 
by Auto_tac;  | 
|
2295  | 
by (rotate_tac 4 1);  | 
|
2296  | 
by (dres_inst_tac [("x","f(x) + xa")] spec 1);
 | 
|
2297  | 
by Auto_tac;  | 
|
2298  | 
by (dtac sym 1 THEN Auto_tac);  | 
|
2299  | 
by (arith_tac 1);  | 
|
| 11383 | 2300  | 
qed "isCont_inverse_function";  | 
| 10751 | 2301  |