12196
|
1 |
(* Title : Transcendental.ML
|
|
2 |
Author : Jacques D. Fleuriot
|
|
3 |
Copyright : 1998,1999 University of Cambridge
|
|
4 |
1999 University of Edinburgh
|
|
5 |
Description : Power Series
|
|
6 |
*)
|
|
7 |
|
|
8 |
Goalw [root_def] "root (Suc n) 0 = 0";
|
|
9 |
by (safe_tac (claset() addSIs [some_equality,realpow_zero]
|
|
10 |
addSEs [realpow_zero_zero]));
|
|
11 |
qed "real_root_zero";
|
|
12 |
Addsimps [real_root_zero];
|
|
13 |
|
|
14 |
Goalw [root_def]
|
|
15 |
"0 < x ==> (root(Suc n) x) ^ (Suc n) = x";
|
|
16 |
by (dres_inst_tac [("n","n")] realpow_pos_nth2 1);
|
|
17 |
by (auto_tac (claset() addIs [someI2],simpset()));
|
|
18 |
qed "real_root_pow_pos";
|
|
19 |
|
|
20 |
Goal "0 <= x ==> (root(Suc n) x) ^ (Suc n) = x";
|
|
21 |
by (auto_tac (claset() addSDs [real_le_imp_less_or_eq]
|
|
22 |
addDs [real_root_pow_pos],simpset()));
|
|
23 |
qed "real_root_pow_pos2";
|
|
24 |
|
|
25 |
Goalw [root_def]
|
|
26 |
"0 < x ==> root(Suc n) (x ^ (Suc n)) = x";
|
|
27 |
by (rtac some_equality 1);
|
|
28 |
by (forw_inst_tac [("n","n")] realpow_gt_zero 2);
|
|
29 |
by (auto_tac (claset(),simpset() addsimps [real_0_less_mult_iff]));
|
|
30 |
by (res_inst_tac [("R1.0","u"),("R2.0","x")] real_linear_less2 1);
|
|
31 |
by (dres_inst_tac [("n3","n"),("x","u")]
|
|
32 |
(zero_less_Suc RSN (3,conjI RSN (2,conjI RS realpow_less))) 1);
|
|
33 |
by (dres_inst_tac [("n3","n"),("x","x")]
|
|
34 |
(zero_less_Suc RSN (3,conjI RSN (2,conjI RS realpow_less))) 4);
|
|
35 |
by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
|
|
36 |
qed "real_root_pos";
|
|
37 |
|
|
38 |
Goal "0 <= x ==> root(Suc n) (x ^ (Suc n)) = x";
|
|
39 |
by (auto_tac (claset() addSDs [real_le_imp_less_or_eq,
|
|
40 |
real_root_pos],simpset()));
|
|
41 |
qed "real_root_pos2";
|
|
42 |
|
|
43 |
Goalw [root_def]
|
|
44 |
"0 < x ==> 0 <= root(Suc n) x";
|
|
45 |
by (dres_inst_tac [("n","n")] realpow_pos_nth2 1);
|
|
46 |
by (Safe_tac THEN rtac someI2 1);
|
|
47 |
by (auto_tac (claset() addSIs [order_less_imp_le]
|
|
48 |
addDs [realpow_gt_zero],simpset() addsimps [real_0_less_mult_iff]));
|
|
49 |
qed "real_root_pos_pos";
|
|
50 |
|
|
51 |
Goal "0 <= x ==> 0 <= root(Suc n) x";
|
|
52 |
by (auto_tac (claset() addSDs [real_le_imp_less_or_eq]
|
|
53 |
addDs [real_root_pos_pos],simpset()));
|
|
54 |
qed "real_root_pos_pos_le";
|
|
55 |
|
|
56 |
Goalw [root_def] "root (Suc n) 1 = 1";
|
|
57 |
by (rtac some_equality 1);
|
|
58 |
by Auto_tac;
|
|
59 |
by (rtac ccontr 1);
|
|
60 |
by (res_inst_tac [("R1.0","u"),("R2.0","1")] real_linear_less2 1);
|
|
61 |
by (dres_inst_tac [("n","n")] realpow_Suc_less_one 1);
|
|
62 |
by (dres_inst_tac [("n","n")] realpow_Suc_gt_one 4);
|
|
63 |
by (auto_tac (claset(),simpset() addsimps [real_less_not_refl]));
|
|
64 |
qed "real_root_one";
|
|
65 |
Addsimps [real_root_one];
|
|
66 |
|
|
67 |
(*----------------------------------------------------------------------*)
|
|
68 |
(* Square root *)
|
|
69 |
(*----------------------------------------------------------------------*)
|
|
70 |
|
|
71 |
(*lcp: needed now because 2 is a binary numeral!*)
|
|
72 |
Goal "root 2 = root (Suc (Suc 0))";
|
|
73 |
by (simp_tac (simpset() delsimps [numeral_0_eq_0, numeral_1_eq_1]
|
|
74 |
addsimps [numeral_0_eq_0 RS sym]) 1);
|
|
75 |
qed "root_2_eq";
|
|
76 |
Addsimps [root_2_eq];
|
|
77 |
|
|
78 |
Goalw [sqrt_def] "sqrt 0 = 0";
|
|
79 |
by (Auto_tac);
|
|
80 |
qed "real_sqrt_zero";
|
|
81 |
Addsimps [real_sqrt_zero];
|
|
82 |
|
|
83 |
Goalw [sqrt_def] "sqrt 1 = 1";
|
|
84 |
by (Auto_tac);
|
|
85 |
qed "real_sqrt_one";
|
|
86 |
Addsimps [real_sqrt_one];
|
|
87 |
|
|
88 |
Goalw [sqrt_def] "(sqrt(x) ^ 2 = x) = (0 <= x)";
|
|
89 |
by (Step_tac 1);
|
|
90 |
by (cut_inst_tac [("r","root 2 x")] realpow_two_le 1);
|
|
91 |
by (stac numeral_2_eq_2 2);
|
|
92 |
by (rtac real_root_pow_pos2 2);
|
|
93 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
94 |
qed "real_sqrt_pow2_iff";
|
|
95 |
Addsimps [real_sqrt_pow2_iff];
|
|
96 |
|
|
97 |
|
|
98 |
Addsimps [realpow_two_le_add_order RS (real_sqrt_pow2_iff RS iffD2)];
|
|
99 |
Addsimps [simplify (simpset()) (realpow_two_le_add_order RS
|
|
100 |
(real_sqrt_pow2_iff RS iffD2))];
|
|
101 |
|
|
102 |
Goalw [sqrt_def] "0 < x ==> sqrt(x) ^ 2 = x";
|
|
103 |
by (stac numeral_2_eq_2 1);
|
|
104 |
by (etac real_root_pow_pos 1);
|
|
105 |
qed "real_sqrt_gt_zero_pow2";
|
|
106 |
|
|
107 |
Goal "(sqrt(abs(x)) ^ 2 = abs x)";
|
|
108 |
by (rtac (real_sqrt_pow2_iff RS iffD2) 1);
|
|
109 |
by (arith_tac 1);
|
|
110 |
qed "real_sqrt_abs_abs";
|
|
111 |
Addsimps [real_sqrt_abs_abs];
|
|
112 |
|
|
113 |
Goalw [sqrt_def]
|
|
114 |
"0 <= x ==> sqrt(x) ^ 2 = sqrt(x ^ 2)";
|
|
115 |
by (stac numeral_2_eq_2 1);
|
|
116 |
by (auto_tac (claset() addIs [real_root_pow_pos2
|
|
117 |
RS ssubst, real_root_pos2 RS ssubst],
|
|
118 |
simpset() delsimps [realpow_Suc]));
|
|
119 |
qed "real_pow_sqrt_eq_sqrt_pow";
|
|
120 |
|
|
121 |
Goal "0 <= x ==> sqrt(x) ^ 2 = sqrt(abs(x ^ 2))";
|
|
122 |
by (asm_full_simp_tac (simpset() addsimps [real_pow_sqrt_eq_sqrt_pow]) 1);
|
|
123 |
by (stac numeral_2_eq_2 1);
|
|
124 |
by (asm_full_simp_tac (simpset()delsimps [realpow_Suc]) 1);
|
|
125 |
qed "real_pow_sqrt_eq_sqrt_abs_pow";
|
|
126 |
|
|
127 |
Goal "0 <= x ==> sqrt(x) ^ 2 = sqrt(abs(x) ^ 2)";
|
|
128 |
by (asm_full_simp_tac (simpset() addsimps [real_pow_sqrt_eq_sqrt_abs_pow]) 1);
|
|
129 |
by (stac numeral_2_eq_2 1);
|
|
130 |
by (asm_full_simp_tac (simpset()delsimps [realpow_Suc]) 1);
|
|
131 |
qed "real_pow_sqrt_eq_sqrt_abs_pow2";
|
|
132 |
|
|
133 |
Goal "0 <= x ==> sqrt(x) ^ 2 = abs(x)";
|
|
134 |
by (rtac (real_sqrt_abs_abs RS subst) 1);
|
|
135 |
by (res_inst_tac [("x1","x")]
|
|
136 |
(real_pow_sqrt_eq_sqrt_abs_pow2 RS ssubst) 1);
|
|
137 |
by (rtac (real_pow_sqrt_eq_sqrt_pow RS sym) 2);
|
|
138 |
by (assume_tac 1 THEN arith_tac 1);
|
|
139 |
qed "real_sqrt_pow_abs";
|
|
140 |
|
|
141 |
Goal "(~ (0::real) < x*x) = (x = 0)";
|
|
142 |
by Auto_tac;
|
|
143 |
by (rtac ccontr 1);
|
|
144 |
by (cut_inst_tac [("R1.0","x"),("R2.0","0")] real_linear 1);
|
|
145 |
by Auto_tac;
|
|
146 |
by (ftac (real_mult_order) 2);
|
|
147 |
by (asm_full_simp_tac (simpset() addsimps [real_0_less_mult_iff]) 1);
|
|
148 |
by Auto_tac;
|
|
149 |
qed "not_real_square_gt_zero";
|
|
150 |
Addsimps [not_real_square_gt_zero];
|
|
151 |
|
|
152 |
|
|
153 |
(* proof used to be simpler *)
|
|
154 |
Goalw [sqrt_def,root_def]
|
|
155 |
"[| 0 < x; 0 < y |] ==>sqrt(x*y) = sqrt(x) * sqrt(y)";
|
|
156 |
by (dres_inst_tac [("n","1")] realpow_pos_nth2 1);
|
|
157 |
by (dres_inst_tac [("n","1")] realpow_pos_nth2 1);
|
|
158 |
by (asm_full_simp_tac (simpset() delsimps [realpow_Suc]
|
|
159 |
addsimps [numeral_2_eq_2]) 1);
|
|
160 |
by (Step_tac 1);
|
|
161 |
by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2);
|
|
162 |
by (Asm_full_simp_tac 1 THEN Asm_full_simp_tac 1);
|
|
163 |
by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2);
|
|
164 |
by (Asm_full_simp_tac 1 THEN Asm_full_simp_tac 1);
|
|
165 |
by (res_inst_tac [("a","xa * x")] someI2 1);
|
|
166 |
by (auto_tac (claset() addEs [real_less_asym],
|
|
167 |
simpset() addsimps real_mult_ac@[realpow_mult RS sym,realpow_two_disj,
|
|
168 |
realpow_gt_zero, real_mult_order] delsimps [realpow_Suc]));
|
|
169 |
qed "real_sqrt_mult_distrib";
|
|
170 |
|
|
171 |
Goal "[|0<=x; 0<=y |] ==> sqrt(x*y) = sqrt(x) * sqrt(y)";
|
|
172 |
by (auto_tac (claset() addIs [ real_sqrt_mult_distrib],
|
|
173 |
simpset() addsimps [real_le_less]));
|
|
174 |
qed "real_sqrt_mult_distrib2";
|
|
175 |
|
|
176 |
Goal "(r * r = 0) = (r = (0::real))";
|
|
177 |
by Auto_tac;
|
|
178 |
qed "real_mult_self_eq_zero_iff";
|
|
179 |
Addsimps [real_mult_self_eq_zero_iff];
|
|
180 |
|
|
181 |
Goalw [sqrt_def,root_def] "0 < x ==> 0 < sqrt(x)";
|
|
182 |
by (stac numeral_2_eq_2 1);
|
|
183 |
by (dtac realpow_pos_nth2 1 THEN Step_tac 1);
|
|
184 |
by (rtac someI2 1 THEN Step_tac 1 THEN Blast_tac 2);
|
|
185 |
by Auto_tac;
|
|
186 |
qed "real_sqrt_gt_zero";
|
|
187 |
|
|
188 |
Goal "0 <= x ==> 0 <= sqrt(x)";
|
|
189 |
by (auto_tac (claset() addIs [real_sqrt_gt_zero],
|
|
190 |
simpset() addsimps [real_le_less]));
|
|
191 |
qed "real_sqrt_ge_zero";
|
|
192 |
|
|
193 |
Goal "0 <= sqrt (x ^ 2 + y ^ 2)";
|
|
194 |
by (auto_tac (claset() addSIs [real_sqrt_ge_zero],simpset()));
|
|
195 |
qed "real_sqrt_sum_squares_ge_zero";
|
|
196 |
Addsimps [real_sqrt_sum_squares_ge_zero];
|
|
197 |
|
|
198 |
Goal "0 <= sqrt ((x ^ 2 + y ^ 2)*(xa ^ 2 + ya ^ 2))";
|
|
199 |
by (auto_tac (claset() addSIs [real_sqrt_ge_zero],simpset()
|
|
200 |
addsimps [real_0_le_mult_iff]));
|
|
201 |
qed "real_sqrt_sum_squares_mult_ge_zero";
|
|
202 |
Addsimps [real_sqrt_sum_squares_mult_ge_zero];
|
|
203 |
|
|
204 |
Goal "sqrt ((x ^ 2 + y ^ 2) * (xa ^ 2 + ya ^ 2)) ^ 2 = \
|
|
205 |
\ (x ^ 2 + y ^ 2) * (xa ^ 2 + ya ^ 2)";
|
|
206 |
by (auto_tac (claset(),simpset() addsimps [real_sqrt_pow2_iff,
|
|
207 |
real_0_le_mult_iff] delsimps [realpow_Suc]));
|
|
208 |
qed "real_sqrt_sum_squares_mult_squared_eq";
|
|
209 |
Addsimps [real_sqrt_sum_squares_mult_squared_eq];
|
|
210 |
|
|
211 |
Goal "sqrt(x ^ 2) = abs(x)";
|
|
212 |
by (rtac (abs_realpow_two RS subst) 1);
|
|
213 |
by (rtac (real_sqrt_abs_abs RS subst) 1);
|
12486
|
214 |
by (stac real_pow_sqrt_eq_sqrt_pow 1);
|
13097
|
215 |
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, abs_mult]));
|
12196
|
216 |
qed "real_sqrt_abs";
|
|
217 |
Addsimps [real_sqrt_abs];
|
|
218 |
|
|
219 |
Goal "sqrt(x*x) = abs(x)";
|
|
220 |
by (rtac (realpow_two RS subst) 1);
|
|
221 |
by (stac (numeral_2_eq_2 RS sym) 1);
|
|
222 |
by (rtac real_sqrt_abs 1);
|
|
223 |
qed "real_sqrt_abs2";
|
|
224 |
Addsimps [real_sqrt_abs2];
|
|
225 |
|
|
226 |
Goal "0 < x ==> 0 < sqrt(x) ^ 2";
|
|
227 |
by (asm_full_simp_tac (simpset() addsimps [real_sqrt_gt_zero_pow2]) 1);
|
|
228 |
qed "real_sqrt_pow2_gt_zero";
|
|
229 |
|
|
230 |
Goal "0 < x ==> sqrt x ~= 0";
|
12486
|
231 |
by (ftac real_sqrt_pow2_gt_zero 1);
|
12196
|
232 |
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2, real_less_not_refl]));
|
|
233 |
qed "real_sqrt_not_eq_zero";
|
|
234 |
|
|
235 |
Goal "0 < x ==> inverse (sqrt(x)) ^ 2 = inverse x";
|
12486
|
236 |
by (ftac real_sqrt_not_eq_zero 1);
|
12196
|
237 |
by (cut_inst_tac [("n1","2"),("r1","sqrt x")] (realpow_inverse RS sym) 1);
|
|
238 |
by (auto_tac (claset() addDs [real_sqrt_gt_zero_pow2],simpset()));
|
|
239 |
qed "real_inv_sqrt_pow2";
|
|
240 |
|
|
241 |
Goal "[| 0 <= x; sqrt(x) = 0|] ==> x = 0";
|
|
242 |
by (dtac real_le_imp_less_or_eq 1);
|
|
243 |
by (auto_tac (claset() addDs [real_sqrt_not_eq_zero],simpset()));
|
|
244 |
qed "real_sqrt_eq_zero_cancel";
|
|
245 |
|
|
246 |
Goal "0 <= x ==> ((sqrt x = 0) = (x = 0))";
|
|
247 |
by (auto_tac (claset(),simpset() addsimps [real_sqrt_eq_zero_cancel]));
|
|
248 |
qed "real_sqrt_eq_zero_cancel_iff";
|
|
249 |
Addsimps [real_sqrt_eq_zero_cancel_iff];
|
|
250 |
|
|
251 |
Goal "x <= sqrt(x ^ 2 + y ^ 2)";
|
|
252 |
by (subgoal_tac "x <= 0 | 0 <= x" 1);
|
|
253 |
by (Step_tac 1);
|
|
254 |
by (rtac real_le_trans 1);
|
|
255 |
by (auto_tac (claset(),simpset() delsimps [realpow_Suc]));
|
|
256 |
by (res_inst_tac [("n","1")] realpow_increasing 1);
|
|
257 |
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2 RS sym]
|
|
258 |
delsimps [realpow_Suc]));
|
|
259 |
qed "real_sqrt_sum_squares_ge1";
|
|
260 |
Addsimps [real_sqrt_sum_squares_ge1];
|
|
261 |
|
|
262 |
Goal "y <= sqrt(z ^ 2 + y ^ 2)";
|
|
263 |
by (simp_tac (simpset() addsimps [real_add_commute]
|
|
264 |
delsimps [realpow_Suc]) 1);
|
|
265 |
qed "real_sqrt_sum_squares_ge2";
|
|
266 |
Addsimps [real_sqrt_sum_squares_ge2];
|
|
267 |
|
|
268 |
Goal "1 <= x ==> 1 <= sqrt x";
|
|
269 |
by (res_inst_tac [("n","1")] realpow_increasing 1);
|
|
270 |
by (auto_tac (claset(),simpset() addsimps [numeral_2_eq_2 RS sym, real_sqrt_gt_zero_pow2,
|
|
271 |
real_sqrt_ge_zero] delsimps [realpow_Suc]));
|
|
272 |
qed "real_sqrt_ge_one";
|
|
273 |
|
|
274 |
(*-------------------------------------------------------------------------*)
|
|
275 |
(* Exponential function *)
|
|
276 |
(*-------------------------------------------------------------------------*)
|
|
277 |
|
|
278 |
Goal "summable (%n. inverse (real (fact n)) * x ^ n)";
|
|
279 |
by (cut_facts_tac [real_zero_less_one RS real_dense] 1);
|
|
280 |
by (Step_tac 1);
|
|
281 |
by (cut_inst_tac [("x","r")] reals_Archimedean3 1);
|
|
282 |
by Auto_tac;
|
|
283 |
by (dres_inst_tac [("x","abs x")] spec 1 THEN Step_tac 1);
|
|
284 |
by (res_inst_tac [("N","n"),("c","r")] ratio_test 1);
|
|
285 |
by (auto_tac (claset(),simpset() addsimps [abs_mult,real_mult_assoc RS sym]
|
|
286 |
delsimps [fact_Suc]));
|
|
287 |
by (rtac real_mult_le_le_mono2 1);
|
|
288 |
by (res_inst_tac [("w1","abs x")] (real_mult_commute RS ssubst) 2);
|
12486
|
289 |
by (stac fact_Suc 2);
|
|
290 |
by (stac real_of_nat_mult 2);
|
13097
|
291 |
by (auto_tac (claset(),simpset() addsimps [abs_mult,real_inverse_distrib]));
|
12196
|
292 |
by (auto_tac (claset(), simpset() addsimps
|
13097
|
293 |
[real_mult_assoc RS sym, abs_eqI2, real_inverse_gt_0]));
|
12196
|
294 |
by (rtac (CLAIM "x < (y::real) ==> x <= y") 1);
|
|
295 |
by (res_inst_tac [("z1","real (Suc na)")] (real_mult_less_iff1
|
|
296 |
RS iffD1) 1);
|
|
297 |
by (auto_tac (claset(),simpset() addsimps [real_not_refl2 RS not_sym,
|
|
298 |
real_mult_assoc,abs_inverse]));
|
|
299 |
by (rtac real_less_trans 1);
|
|
300 |
by (auto_tac (claset(),simpset() addsimps real_mult_ac));
|
|
301 |
qed "summable_exp";
|
|
302 |
|
|
303 |
Addsimps [real_of_nat_fact_gt_zero,
|
|
304 |
real_of_nat_fact_ge_zero,inv_real_of_nat_fact_gt_zero,
|
|
305 |
inv_real_of_nat_fact_ge_zero];
|
|
306 |
|
|
307 |
Goalw [real_divide_def]
|
|
308 |
"summable (%n. \
|
|
309 |
\ (if even n then 0 \
|
|
310 |
\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * \
|
|
311 |
\ x ^ n)";
|
|
312 |
by (res_inst_tac [("g","(%n. inverse (real (fact n)) * abs(x) ^ n)")]
|
|
313 |
summable_comparison_test 1);
|
|
314 |
by (rtac summable_exp 2);
|
|
315 |
by (res_inst_tac [("x","0")] exI 1);
|
12330
|
316 |
by (auto_tac (claset(), simpset() addsimps [realpow_abs RS sym,
|
13097
|
317 |
abs_mult,real_0_le_mult_iff]));
|
12196
|
318 |
by (auto_tac (claset() addIs [real_mult_le_le_mono2],
|
13097
|
319 |
simpset() addsimps [real_inverse_gt_0,abs_eqI2]));
|
12196
|
320 |
qed "summable_sin";
|
|
321 |
|
|
322 |
Goalw [real_divide_def]
|
|
323 |
"summable (%n. \
|
|
324 |
\ (if even n then \
|
|
325 |
\ (- 1) ^ (n div 2)/(real (fact n)) else 0) * x ^ n)";
|
|
326 |
by (res_inst_tac [("g","(%n. inverse (real (fact n)) * abs(x) ^ n)")]
|
|
327 |
summable_comparison_test 1);
|
|
328 |
by (rtac summable_exp 2);
|
|
329 |
by (res_inst_tac [("x","0")] exI 1);
|
13097
|
330 |
by (auto_tac (claset(), simpset() addsimps [realpow_abs RS sym,abs_mult,
|
12196
|
331 |
real_0_le_mult_iff]));
|
|
332 |
by (auto_tac (claset() addSIs [real_mult_le_le_mono2],
|
13097
|
333 |
simpset() addsimps [real_inverse_gt_0,abs_eqI2]));
|
12196
|
334 |
qed "summable_cos";
|
|
335 |
|
|
336 |
Goal "(if even n then 0 \
|
|
337 |
\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * 0 ^ n = 0";
|
|
338 |
by (induct_tac "n" 1);
|
|
339 |
by (Auto_tac);
|
|
340 |
val lemma_STAR_sin = result();
|
|
341 |
Addsimps [lemma_STAR_sin];
|
|
342 |
|
|
343 |
Goal "0 < n --> \
|
|
344 |
\ (- 1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0";
|
|
345 |
by (induct_tac "n" 1);
|
|
346 |
by (Auto_tac);
|
|
347 |
val lemma_STAR_cos = result();
|
|
348 |
Addsimps [lemma_STAR_cos];
|
|
349 |
|
|
350 |
Goal "0 < n --> \
|
|
351 |
\ (-1) ^ (n div 2)/(real (fact n)) * 0 ^ n = 0";
|
|
352 |
by (induct_tac "n" 1);
|
|
353 |
by (Auto_tac);
|
|
354 |
val lemma_STAR_cos1 = result();
|
|
355 |
Addsimps [lemma_STAR_cos1];
|
|
356 |
|
|
357 |
Goal "sumr 1 n (%n. if even n \
|
|
358 |
\ then (- 1) ^ (n div 2)/(real (fact n)) * \
|
|
359 |
\ 0 ^ n \
|
|
360 |
\ else 0) = 0";
|
|
361 |
by (induct_tac "n" 1);
|
|
362 |
by (case_tac "n" 2);
|
|
363 |
by (Auto_tac);
|
|
364 |
val lemma_STAR_cos2 = result();
|
|
365 |
Addsimps [lemma_STAR_cos2];
|
|
366 |
|
|
367 |
Goalw [exp_def] "(%n. inverse (real (fact n)) * x ^ n) sums exp(x)";
|
|
368 |
by (rtac (summable_exp RS summable_sums) 1);
|
|
369 |
qed "exp_converges";
|
|
370 |
|
|
371 |
Goalw [sin_def]
|
|
372 |
"(%n. (if even n then 0 \
|
|
373 |
\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * \
|
|
374 |
\ x ^ n) sums sin(x)";
|
|
375 |
by (rtac (summable_sin RS summable_sums) 1);
|
|
376 |
qed "sin_converges";
|
|
377 |
|
|
378 |
Goalw [cos_def]
|
|
379 |
"(%n. (if even n then \
|
|
380 |
\ (- 1) ^ (n div 2)/(real (fact n)) \
|
|
381 |
\ else 0) * x ^ n) sums cos(x)";
|
|
382 |
by (rtac (summable_cos RS summable_sums) 1);
|
|
383 |
qed "cos_converges";
|
|
384 |
|
|
385 |
Goal "p <= n --> y ^ (Suc n - p) = ((y::real) ^ (n - p)) * y";
|
|
386 |
by (induct_tac "n" 1 THEN Auto_tac);
|
|
387 |
by (subgoal_tac "p = Suc n" 1);
|
|
388 |
by (Asm_simp_tac 1 THEN Auto_tac);
|
|
389 |
by (dtac sym 1 THEN asm_full_simp_tac (simpset() addsimps
|
|
390 |
[Suc_diff_le,real_mult_commute,realpow_Suc RS sym]
|
|
391 |
delsimps [realpow_Suc]) 1);
|
|
392 |
qed_spec_mp "lemma_realpow_diff";
|
|
393 |
|
|
394 |
(*--------------------------------------------------------------------------*)
|
|
395 |
(* Properties of power series *)
|
|
396 |
(*--------------------------------------------------------------------------*)
|
|
397 |
|
|
398 |
Goal "sumr 0 (Suc n) (%p. (x ^ p) * y ^ ((Suc n) - p)) = \
|
|
399 |
\ y * sumr 0 (Suc n) (%p. (x ^ p) * (y ^ (n - p)))";
|
|
400 |
by (auto_tac (claset(),simpset() addsimps [sumr_mult] delsimps [sumr_Suc]));
|
|
401 |
by (rtac sumr_subst 1);
|
|
402 |
by (strip_tac 1);
|
12486
|
403 |
by (stac lemma_realpow_diff 1);
|
12196
|
404 |
by (auto_tac (claset(),simpset() addsimps real_mult_ac));
|
|
405 |
qed "lemma_realpow_diff_sumr";
|
|
406 |
|
|
407 |
Goal "x ^ (Suc n) - y ^ (Suc n) = \
|
|
408 |
\ (x - y) * sumr 0 (Suc n) (%p. (x ^ p) * (y ^(n - p)))";
|
|
409 |
by (induct_tac "n" 1);
|
|
410 |
by (auto_tac (claset(),simpset() delsimps [sumr_Suc]));
|
12486
|
411 |
by (stac sumr_Suc 1);
|
12196
|
412 |
by (dtac sym 1);
|
|
413 |
by (auto_tac (claset(),simpset() addsimps [lemma_realpow_diff_sumr,
|
|
414 |
real_add_mult_distrib2,real_diff_def] @
|
|
415 |
real_mult_ac delsimps [sumr_Suc]));
|
|
416 |
qed "lemma_realpow_diff_sumr2";
|
|
417 |
|
|
418 |
Goal "sumr 0 (Suc n) (%p. (x ^ p) * (y ^ (n - p))) = \
|
|
419 |
\ sumr 0 (Suc n) (%p. (x ^ (n - p)) * (y ^ p))";
|
|
420 |
by (case_tac "x = y" 1);
|
|
421 |
by (auto_tac (claset(),simpset() addsimps [real_mult_commute,
|
|
422 |
realpow_add RS sym] delsimps [sumr_Suc]));
|
|
423 |
by (res_inst_tac [("c1","x - y")] (real_mult_left_cancel RS iffD1) 1);
|
|
424 |
by (rtac (real_minus_minus RS subst) 2);
|
12486
|
425 |
by (stac real_minus_mult_eq1 2);
|
12196
|
426 |
by (auto_tac (claset(),simpset() addsimps [lemma_realpow_diff_sumr2
|
|
427 |
RS sym] delsimps [sumr_Suc]));
|
|
428 |
qed "lemma_realpow_rev_sumr";
|
|
429 |
|
|
430 |
(* ------------------------------------------------------------------------ *)
|
|
431 |
(* Power series has a `circle` of convergence, *)
|
|
432 |
(* i.e. if it sums for x, then it sums absolutely for z with |z| < |x|. *)
|
|
433 |
(* ------------------------------------------------------------------------ *)
|
|
434 |
|
|
435 |
Goalw [real_divide_def] "1/(x::real) = inverse x";
|
|
436 |
by (Simp_tac 1);
|
|
437 |
qed "real_divide_eq_inverse";
|
|
438 |
|
|
439 |
Goal "[| summable (%n. f(n) * (x ^ n)); abs(z) < abs(x) |] \
|
|
440 |
\ ==> summable (%n. abs(f(n)) * (z ^ n))";
|
|
441 |
by (dtac summable_LIMSEQ_zero 1);
|
|
442 |
by (dtac convergentI 1);
|
|
443 |
by (asm_full_simp_tac (simpset() addsimps [Cauchy_convergent_iff RS sym]) 1);
|
|
444 |
by (dtac Cauchy_Bseq 1);
|
|
445 |
by (asm_full_simp_tac (simpset() addsimps [Bseq_def]) 1);
|
|
446 |
by (Step_tac 1);
|
|
447 |
by (res_inst_tac [("g","%n. K * abs(z ^ n) * inverse (abs(x ^ n))")]
|
|
448 |
summable_comparison_test 1);
|
|
449 |
by (res_inst_tac [("x","0")] exI 1 THEN Step_tac 1);
|
|
450 |
by (subgoal_tac "0 < abs (x ^ n)" 1);
|
|
451 |
by (res_inst_tac [("z","abs (x ^ n)")] (CLAIM_SIMP
|
|
452 |
"[| (0::real) <z; x*z<=y*z |] ==> x<=y" [real_mult_le_cancel1]) 1);
|
|
453 |
by (auto_tac (claset(),
|
12330
|
454 |
simpset() addsimps [real_mult_assoc,realpow_abs]));
|
12196
|
455 |
by (dres_inst_tac [("x","0")] spec 2 THEN Force_tac 2);
|
13097
|
456 |
by (auto_tac (claset(),simpset() addsimps [abs_mult,realpow_abs] @ real_mult_ac));
|
12196
|
457 |
by (res_inst_tac [("x2","z ^ n")] (abs_ge_zero RS real_le_imp_less_or_eq
|
|
458 |
RS disjE) 1 THEN dtac sym 2);
|
|
459 |
by (auto_tac (claset() addSIs [real_mult_le_le_mono2],
|
|
460 |
simpset() addsimps [real_mult_assoc RS sym,
|
12330
|
461 |
realpow_abs,summable_def]));
|
12196
|
462 |
by (res_inst_tac [("x","K * inverse(1 - (abs(z) * inverse(abs x)))")] exI 1);
|
|
463 |
by (auto_tac (claset() addSIs [sums_mult],simpset() addsimps [real_mult_assoc]));
|
|
464 |
by (subgoal_tac
|
|
465 |
"abs(z ^ n) * inverse(abs x ^ n) = (abs(z) * inverse(abs x)) ^ n" 1);
|
12330
|
466 |
by (auto_tac (claset(),simpset() addsimps [realpow_abs RS sym]));
|
12196
|
467 |
by (subgoal_tac "x ~= 0" 1);
|
|
468 |
by (subgoal_tac "x ~= 0" 3);
|
|
469 |
by (auto_tac (claset(),simpset() addsimps
|
|
470 |
[abs_inverse RS sym,realpow_not_zero,abs_mult
|
|
471 |
RS sym,realpow_inverse,realpow_mult RS sym]));
|
|
472 |
by (auto_tac (claset() addSIs [geometric_sums],simpset() addsimps
|
12330
|
473 |
[realpow_abs,real_divide_eq_inverse RS sym]));
|
12196
|
474 |
by (res_inst_tac [("z","abs(x)")] (CLAIM_SIMP
|
|
475 |
"[|(0::real)<z; x*z<y*z |] ==> x<y" [real_mult_less_cancel1]) 1);
|
13097
|
476 |
by (auto_tac (claset(),simpset() addsimps [abs_mult RS sym,real_mult_assoc]));
|
12196
|
477 |
qed "powser_insidea";
|
|
478 |
|
|
479 |
Goal "[| summable (%n. f(n) * (x ^ n)); abs(z) < abs(x) |] \
|
|
480 |
\ ==> summable (%n. f(n) * (z ^ n))";
|
|
481 |
by (dres_inst_tac [("z","abs z")] powser_insidea 1);
|
|
482 |
by (auto_tac (claset() addIs [summable_rabs_cancel],
|
12330
|
483 |
simpset() addsimps [realpow_abs RS sym,abs_mult RS sym]));
|
12196
|
484 |
qed "powser_inside";
|
|
485 |
|
|
486 |
(* ------------------------------------------------------------------------ *)
|
|
487 |
(* Differentiation of power series *)
|
|
488 |
(* ------------------------------------------------------------------------ *)
|
|
489 |
|
|
490 |
(* Lemma about distributing negation over it *)
|
|
491 |
Goalw [diffs_def] "diffs (%n. - c n) = (%n. - diffs c n)";
|
|
492 |
by Auto_tac;
|
|
493 |
qed "diffs_minus";
|
|
494 |
|
|
495 |
(* ------------------------------------------------------------------------ *)
|
|
496 |
(* Show that we can shift the terms down one *)
|
|
497 |
(* ------------------------------------------------------------------------ *)
|
|
498 |
|
|
499 |
Goal "sumr 0 n (%n. (diffs c)(n) * (x ^ n)) = \
|
|
500 |
\ sumr 0 n (%n. real n * c(n) * (x ^ (n - Suc 0))) + \
|
|
501 |
\ (real n * c(n) * x ^ (n - Suc 0))";
|
|
502 |
by (induct_tac "n" 1);
|
|
503 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
|
|
504 |
real_add_assoc RS sym,diffs_def]));
|
|
505 |
qed "lemma_diffs";
|
|
506 |
|
|
507 |
Goal "sumr 0 n (%n. real n * c(n) * (x ^ (n - Suc 0))) = \
|
|
508 |
\ sumr 0 n (%n. (diffs c)(n) * (x ^ n)) - \
|
|
509 |
\ (real n * c(n) * x ^ (n - Suc 0))";
|
|
510 |
by (auto_tac (claset(),simpset() addsimps [lemma_diffs]));
|
|
511 |
qed "lemma_diffs2";
|
|
512 |
|
|
513 |
Goal "summable (%n. (diffs c)(n) * (x ^ n)) ==> \
|
|
514 |
\ (%n. real n * c(n) * (x ^ (n - Suc 0))) sums \
|
|
515 |
\ (suminf(%n. (diffs c)(n) * (x ^ n)))";
|
|
516 |
by (ftac summable_LIMSEQ_zero 1);
|
|
517 |
by (subgoal_tac "(%n. real n * c(n) * (x ^ (n - Suc 0))) ----> 0" 1);
|
|
518 |
by (rtac LIMSEQ_imp_Suc 2);
|
|
519 |
by (dtac summable_sums 1);
|
|
520 |
by (auto_tac (claset(),simpset() addsimps [sums_def]));
|
|
521 |
by (thin_tac "(%n. diffs c n * x ^ n) ----> 0" 1);
|
|
522 |
by (rotate_tac 1 1);
|
|
523 |
by (dtac LIMSEQ_diff 1);
|
|
524 |
by (auto_tac (claset(),simpset() addsimps [lemma_diffs2 RS sym,
|
|
525 |
symmetric diffs_def]));
|
|
526 |
by (asm_full_simp_tac (simpset() addsimps [diffs_def]) 1);
|
|
527 |
qed "diffs_equiv";
|
|
528 |
|
|
529 |
(* -------------------------------------------------------------------------*)
|
|
530 |
(* Term-by-term differentiability of power series *)
|
|
531 |
(* -------------------------------------------------------------------------*)
|
|
532 |
|
|
533 |
Goal "sumr 0 m (%p. (((z + h) ^ (m - p)) * (z ^ p)) - (z ^ m)) = \
|
|
534 |
\ sumr 0 m (%p. (z ^ p) * \
|
|
535 |
\ (((z + h) ^ (m - p)) - (z ^ (m - p))))";
|
|
536 |
by (rtac sumr_subst 1);
|
|
537 |
by (auto_tac (claset(),simpset() addsimps [real_add_mult_distrib2,
|
|
538 |
real_diff_def,realpow_add RS sym]
|
|
539 |
@ real_mult_ac));
|
|
540 |
qed "lemma_termdiff1";
|
|
541 |
|
|
542 |
(* proved elsewhere? *)
|
|
543 |
Goal "m < n --> (EX d. n = m + d + Suc 0)";
|
|
544 |
by (induct_tac "m" 1 THEN Auto_tac);
|
|
545 |
by (case_tac "n" 1);
|
|
546 |
by (case_tac "d" 3);
|
|
547 |
by (Auto_tac);
|
|
548 |
qed_spec_mp "less_add_one";
|
|
549 |
|
|
550 |
Goal " h ~= 0 ==> \
|
|
551 |
\ (((z + h) ^ n) - (z ^ n)) * inverse h - \
|
|
552 |
\ real n * (z ^ (n - Suc 0)) = \
|
|
553 |
\ h * sumr 0 (n - Suc 0) (%p. (z ^ p) * \
|
|
554 |
\ sumr 0 ((n - Suc 0) - p) \
|
|
555 |
\ (%q. ((z + h) ^ q) * (z ^ (((n - 2) - p) - q))))";
|
|
556 |
by (rtac (real_mult_left_cancel RS iffD1) 1 THEN Asm_simp_tac 1);
|
|
557 |
by (asm_full_simp_tac (simpset() addsimps [real_diff_mult_distrib2]
|
|
558 |
@ real_mult_ac) 1);
|
|
559 |
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 1);
|
|
560 |
by (case_tac "n" 1 THEN auto_tac (claset(),simpset()
|
|
561 |
addsimps [lemma_realpow_diff_sumr2,
|
|
562 |
real_diff_mult_distrib2 RS sym,real_mult_assoc]
|
|
563 |
delsimps [realpow_Suc,sumr_Suc]));
|
|
564 |
by (rtac (real_mult_left_cancel RS iffD2) 1);
|
|
565 |
by (auto_tac (claset(),simpset() addsimps [lemma_realpow_rev_sumr]
|
|
566 |
delsimps [sumr_Suc]));
|
|
567 |
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc,sumr_diff_mult_const,
|
|
568 |
real_add_mult_distrib,CLAIM "(a + b) - (c + d) = a - c + b - (d::real)",
|
|
569 |
lemma_termdiff1,sumr_mult]));
|
|
570 |
by (auto_tac (claset() addSIs [sumr_subst],simpset() addsimps
|
|
571 |
[real_diff_def,real_add_assoc]));
|
|
572 |
by (fold_tac [real_diff_def] THEN dtac less_add_one 1);
|
|
573 |
by (auto_tac (claset(),simpset() addsimps [sumr_mult,lemma_realpow_diff_sumr2]
|
|
574 |
@ real_mult_ac delsimps [sumr_Suc,realpow_Suc]));
|
|
575 |
qed "lemma_termdiff2";
|
|
576 |
|
|
577 |
Goal "[| h ~= 0; abs z <= K; abs (z + h) <= K |] \
|
|
578 |
\ ==> abs (((z + h) ^ n - z ^ n) * inverse h - real n * z ^ (n - Suc 0)) \
|
|
579 |
\ <= real n * real (n - Suc 0) * K ^ (n - 2) * abs h";
|
12486
|
580 |
by (stac lemma_termdiff2 1);
|
12196
|
581 |
by (asm_full_simp_tac (simpset() addsimps [abs_mult,real_mult_commute]) 2);
|
|
582 |
by (stac real_mult_commute 2);
|
|
583 |
by (rtac (sumr_rabs RS real_le_trans) 2);
|
|
584 |
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 2);
|
|
585 |
by (rtac (real_mult_commute RS subst) 2);
|
13097
|
586 |
by (auto_tac (claset() addSIs [sumr_bound2],simpset() addsimps [abs_mult]));
|
12196
|
587 |
by (case_tac "n" 1 THEN Auto_tac);
|
|
588 |
by (dtac less_add_one 1);
|
|
589 |
by (auto_tac (claset(),simpset() addsimps [realpow_add,real_add_assoc RS sym,
|
|
590 |
CLAIM_SIMP "(a * b) * c = a * (c * (b::real))" real_mult_ac]
|
|
591 |
delsimps [sumr_Suc]));
|
13097
|
592 |
by (auto_tac (claset() addSIs [real_mult_le_mono],simpset()delsimps [sumr_Suc]));
|
12196
|
593 |
by (auto_tac (claset() addSIs [realpow_le2],simpset() addsimps
|
13097
|
594 |
[realpow_abs] delsimps [sumr_Suc] ));
|
12196
|
595 |
by (res_inst_tac [("j","real (Suc d) * (K ^ d)")] real_le_trans 1);
|
|
596 |
by (subgoal_tac "0 <= K" 2);
|
|
597 |
by (arith_tac 3);
|
|
598 |
by (dres_inst_tac [("n","d")] realpow_ge_zero2 2);
|
13097
|
599 |
by (auto_tac (claset(),simpset() delsimps [sumr_Suc] ));
|
12196
|
600 |
by (rtac (sumr_rabs RS real_le_trans) 1);
|
|
601 |
by (rtac sumr_bound2 1 THEN auto_tac (claset() addSDs [less_add_one]
|
13097
|
602 |
addSIs [real_mult_le_mono],simpset() addsimps [abs_mult, realpow_add]));
|
12196
|
603 |
by (auto_tac (claset() addSIs [realpow_le2,realpow_ge_zero],simpset() addsimps
|
13097
|
604 |
[realpow_abs]));
|
12196
|
605 |
by (ALLGOALS(arith_tac));
|
|
606 |
qed "lemma_termdiff3";
|
|
607 |
|
|
608 |
Goalw [LIM_def]
|
|
609 |
"[| 0 < k; \
|
|
610 |
\ (ALL h. 0 < abs(h) & abs(h) < k --> abs(f h) <= K * abs(h)) |] \
|
|
611 |
\ ==> f -- 0 --> 0";
|
|
612 |
by (Auto_tac);
|
|
613 |
by (subgoal_tac "0 <= K" 1);
|
|
614 |
by (dres_inst_tac [("x","k*inverse 2")] spec 2);
|
|
615 |
by (ftac real_less_half_sum 2);
|
|
616 |
by (dtac real_gt_half_sum 2);
|
|
617 |
by (auto_tac (claset(),simpset() addsimps [abs_eqI2]));
|
|
618 |
by (res_inst_tac [("z","k/2")] (CLAIM_SIMP
|
|
619 |
"[| (0::real) <z; x*z<=y*z |] ==> x<=y" [real_mult_le_cancel1]) 2);
|
|
620 |
by (auto_tac (claset() addIs [abs_ge_zero RS real_le_trans],simpset()));
|
|
621 |
by (dtac real_le_imp_less_or_eq 1);
|
|
622 |
by Auto_tac;
|
|
623 |
by (subgoal_tac "0 < (r * inverse K) * inverse 2" 1);
|
|
624 |
by (REPEAT(rtac (real_mult_order) 2));
|
|
625 |
by (dres_inst_tac [("d1.0","r * inverse K * inverse 2"),("d2.0","k")]
|
|
626 |
real_lbound_gt_zero 1);
|
|
627 |
by (auto_tac (claset(),simpset() addsimps [real_inverse_gt_0,
|
|
628 |
real_0_less_mult_iff]));
|
|
629 |
by (rtac real_le_trans 2 THEN assume_tac 3 THEN Auto_tac);
|
|
630 |
by (res_inst_tac [("x","e")] exI 1 THEN Auto_tac);
|
|
631 |
by (res_inst_tac [("y","K * abs x")] order_le_less_trans 1);
|
|
632 |
by (res_inst_tac [("R2.0","K * e")] real_less_trans 2);
|
|
633 |
by (res_inst_tac [("z","inverse K")] (CLAIM_SIMP
|
|
634 |
"[|(0::real) <z; z*x<z*y |] ==> x<y" [real_mult_less_cancel1]) 3);
|
|
635 |
by (asm_full_simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 4);
|
|
636 |
by Auto_tac;
|
|
637 |
by (auto_tac (claset(),simpset() addsimps real_mult_ac));
|
|
638 |
qed "lemma_termdiff4";
|
|
639 |
|
|
640 |
Goal "[| 0 < k; \
|
|
641 |
\ summable f; \
|
|
642 |
\ ALL h. 0 < abs(h) & abs(h) < k --> \
|
|
643 |
\ (ALL n. abs(g(h) (n::nat)) <= (f(n) * abs(h))) |] \
|
|
644 |
\ ==> (%h. suminf(g h)) -- 0 --> 0";
|
|
645 |
by (dtac summable_sums 1);
|
|
646 |
by (subgoal_tac "ALL h. 0 < abs h & abs h < k --> \
|
|
647 |
\ abs(suminf (g h)) <= suminf f * abs h" 1);
|
|
648 |
by (Auto_tac);
|
|
649 |
by (subgoal_tac "summable (%n. f n * abs h)" 2);
|
|
650 |
by (simp_tac (simpset() addsimps [summable_def]) 3);
|
|
651 |
by (res_inst_tac [("x","suminf f * abs h")] exI 3);
|
|
652 |
by (dres_inst_tac [("c","abs h")] sums_mult 3);
|
|
653 |
by (asm_full_simp_tac (simpset() addsimps real_mult_ac) 3);
|
|
654 |
by (subgoal_tac "summable (%n. abs(g(h::real)(n::nat)))" 2);
|
|
655 |
by (res_inst_tac [("g","%n. f(n::nat) * abs(h)")] summable_comparison_test 3);
|
|
656 |
by (res_inst_tac [("x","0")] exI 3);
|
|
657 |
by Auto_tac;
|
|
658 |
by (res_inst_tac [("j","suminf(%n. abs(g h n))")] real_le_trans 2);
|
|
659 |
by (auto_tac (claset() addIs [summable_rabs,summable_le],simpset() addsimps
|
|
660 |
[sums_summable RS suminf_mult]));
|
|
661 |
by (auto_tac (claset() addSIs [lemma_termdiff4],simpset() addsimps
|
|
662 |
[(sums_summable RS suminf_mult) RS sym]));
|
|
663 |
qed "lemma_termdiff5";
|
|
664 |
|
|
665 |
(* FIXME: Long proof *)
|
|
666 |
Goalw [deriv_def]
|
|
667 |
"[| summable(%n. c(n) * (K ^ n)); \
|
|
668 |
\ summable(%n. (diffs c)(n) * (K ^ n)); \
|
|
669 |
\ summable(%n. (diffs(diffs c))(n) * (K ^ n)); \
|
|
670 |
\ abs(x) < abs(K) |] \
|
|
671 |
\ ==> DERIV (%x. suminf (%n. c(n) * (x ^ n))) x :> \
|
|
672 |
\ suminf (%n. (diffs c)(n) * (x ^ n))";
|
|
673 |
|
|
674 |
by (res_inst_tac [("g","%h. suminf(%n. ((c(n) * ((x + h) ^ n)) - \
|
|
675 |
\ (c(n) * (x ^ n))) * inverse h)")] LIM_trans 1);
|
|
676 |
by (asm_full_simp_tac (simpset() addsimps [LIM_def]) 1);
|
|
677 |
by (Step_tac 1);
|
|
678 |
by (res_inst_tac [("x","abs K - abs x")] exI 1);
|
|
679 |
by (auto_tac (claset(),simpset() addsimps [real_less_diff_eq]));
|
|
680 |
by (dtac (abs_triangle_ineq RS order_le_less_trans) 1);
|
|
681 |
by (res_inst_tac [("y","0")] order_le_less_trans 1);
|
|
682 |
by Auto_tac;
|
|
683 |
by (subgoal_tac "(%n. (c n) * (x ^ n)) sums \
|
|
684 |
\ (suminf(%n. (c n) * (x ^ n))) & \
|
|
685 |
\ (%n. (c n) * ((x + xa) ^ n)) sums \
|
|
686 |
\ (suminf(%n. (c n) * ((x + xa) ^ n)))" 1);
|
|
687 |
by (auto_tac (claset() addSIs [summable_sums],simpset()));
|
|
688 |
by (rtac powser_inside 2 THEN rtac powser_inside 4);
|
|
689 |
by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
|
|
690 |
by (EVERY1[rotate_tac 8, dtac sums_diff, assume_tac]);
|
|
691 |
by (dres_inst_tac [("x","(%n. c n * (xa + x) ^ n - c n * x ^ n)"),
|
|
692 |
("c","inverse xa")] sums_mult 1);
|
|
693 |
by (rtac (sums_unique RS sym) 1);
|
|
694 |
by (asm_full_simp_tac (simpset() addsimps [real_diff_def,
|
|
695 |
real_divide_def] @ real_add_ac @ real_mult_ac) 1);
|
|
696 |
by (rtac LIM_zero_cancel 1);
|
|
697 |
by (res_inst_tac [("g","%h. suminf (%n. c(n) * (((((x + h) ^ n) - \
|
|
698 |
\ (x ^ n)) * inverse h) - (real n * (x ^ (n - Suc 0)))))")] LIM_trans 1);
|
|
699 |
by (asm_full_simp_tac (simpset() addsimps [LIM_def]) 1);
|
|
700 |
by (Step_tac 1);
|
|
701 |
by (res_inst_tac [("x","abs K - abs x")] exI 1);
|
|
702 |
by (auto_tac (claset(),simpset() addsimps [real_less_diff_eq]));
|
|
703 |
by (dtac (abs_triangle_ineq RS order_le_less_trans) 1);
|
|
704 |
by (res_inst_tac [("y","0")] order_le_less_trans 1);
|
|
705 |
by Auto_tac;
|
|
706 |
by (subgoal_tac "summable(%n. (diffs c)(n) * (x ^ n))" 1);
|
|
707 |
by (rtac powser_inside 2);
|
|
708 |
by (Auto_tac);
|
|
709 |
by (dres_inst_tac [("c","c"),("x","x")] diffs_equiv 1);
|
|
710 |
by (ftac sums_unique 1 THEN Auto_tac);
|
|
711 |
by (subgoal_tac "(%n. (c n) * (x ^ n)) sums \
|
|
712 |
\ (suminf(%n. (c n) * (x ^ n))) & \
|
|
713 |
\ (%n. (c n) * ((x + xa) ^ n)) sums \
|
|
714 |
\ (suminf(%n. (c n) * ((x + xa) ^ n)))" 1);
|
|
715 |
by (Step_tac 1);
|
|
716 |
by (auto_tac (claset() addSIs [summable_sums],simpset()));
|
|
717 |
by (rtac powser_inside 2 THEN rtac powser_inside 4);
|
|
718 |
by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
|
|
719 |
by (forw_inst_tac [("x","(%n. c n * (xa + x) ^ n)"),
|
|
720 |
("y","(%n. c n * x ^ n)")] sums_diff 1 THEN assume_tac 1);
|
|
721 |
by (asm_full_simp_tac (simpset() addsimps [[sums_summable,sums_summable]
|
|
722 |
MRS suminf_diff,real_diff_mult_distrib2 RS sym]) 1);
|
|
723 |
by (forw_inst_tac [("x","(%n. c n * ((xa + x) ^ n - x ^ n))"),
|
|
724 |
("c","inverse xa")] sums_mult 1);
|
|
725 |
by (asm_full_simp_tac (simpset() addsimps [sums_summable RS suminf_mult2]) 1);
|
|
726 |
by (forw_inst_tac [("x","(%n. inverse xa * (c n * ((xa + x) ^ n - x ^ n)))"),
|
|
727 |
("y","(%n. real n * c n * x ^ (n - Suc 0))")] sums_diff 1);
|
|
728 |
by (assume_tac 1);
|
|
729 |
by (rtac (ARITH_PROVE "z - y = x ==> - x = (y::real) - z") 1);
|
|
730 |
by (asm_full_simp_tac (simpset() addsimps [[sums_summable,sums_summable]
|
|
731 |
MRS suminf_diff] @ real_add_ac @ real_mult_ac ) 1);
|
|
732 |
by (res_inst_tac [("f","suminf")] arg_cong 1);
|
|
733 |
by (rtac ext 1);
|
|
734 |
by (asm_full_simp_tac (simpset() addsimps [real_diff_def,
|
|
735 |
real_add_mult_distrib2] @ real_add_ac @ real_mult_ac) 1);
|
|
736 |
(* 46 *)
|
|
737 |
by (dtac real_dense 1 THEN Step_tac 1);
|
|
738 |
by (ftac (real_less_sum_gt_zero) 1);
|
|
739 |
by (dres_inst_tac [("f","%n. abs(c n) * real n * \
|
|
740 |
\ real (n - Suc 0) * (r ^ (n - 2))"),
|
|
741 |
("g","%h n. c(n) * (((((x + h) ^ n) - (x ^ n)) * inverse h) - \
|
|
742 |
\ (real n * (x ^ (n - Suc 0))))")] lemma_termdiff5 1);
|
|
743 |
by (auto_tac (claset(),simpset() addsimps [real_add_commute]));
|
|
744 |
by (subgoal_tac "summable(%n. abs(diffs(diffs c) n) * (r ^ n))" 1);
|
|
745 |
by (res_inst_tac [("x","K")] powser_insidea 2 THEN Auto_tac);
|
|
746 |
by (subgoal_tac "abs r = r" 2 THEN Auto_tac);
|
|
747 |
by (res_inst_tac [("j1","abs x")] (real_le_trans RS abs_eqI1) 2);
|
|
748 |
by Auto_tac;
|
|
749 |
by (asm_full_simp_tac (simpset() addsimps [diffs_def,abs_mult,
|
|
750 |
real_mult_assoc RS sym]) 1);
|
|
751 |
by (subgoal_tac "ALL n. real (Suc n) * real (Suc(Suc n)) * \
|
|
752 |
\ abs(c(Suc(Suc n))) * (r ^ n) = diffs(diffs (%n. abs(c n))) n * (r ^ n)" 1);
|
|
753 |
by (dres_inst_tac [("P","summable")]
|
|
754 |
(CLAIM "[|ALL n. f(n) = g(n); P(%n. f n)|] ==> P(%n. g(n))") 1);
|
|
755 |
by (Auto_tac);
|
|
756 |
by (asm_full_simp_tac (simpset() addsimps [diffs_def]) 2
|
|
757 |
THEN asm_full_simp_tac (simpset() addsimps [diffs_def]) 2);
|
|
758 |
by (dtac diffs_equiv 1);
|
|
759 |
by (dtac sums_summable 1);
|
|
760 |
by (asm_full_simp_tac (simpset() addsimps [diffs_def] @ real_mult_ac) 1);
|
|
761 |
by (subgoal_tac "(%n. real n * (real (Suc n) * (abs(c(Suc n)) * \
|
|
762 |
\ (r ^ (n - Suc 0))))) = (%n. diffs(%m. real (m - Suc 0) * \
|
|
763 |
\ abs(c m) * inverse r) n * (r ^ n))" 1);
|
|
764 |
by (Auto_tac);
|
|
765 |
by (rtac ext 2);
|
|
766 |
by (asm_full_simp_tac (simpset() addsimps [diffs_def]) 2);
|
|
767 |
by (case_tac "n" 2);
|
|
768 |
by Auto_tac;
|
|
769 |
(* 69 *)
|
|
770 |
by (dtac (abs_ge_zero RS order_le_less_trans) 2);
|
|
771 |
by (asm_full_simp_tac (simpset() addsimps real_mult_ac) 2);
|
|
772 |
by (dtac diffs_equiv 1);
|
|
773 |
by (dtac sums_summable 1);
|
|
774 |
by (res_inst_tac [("a","summable (%n. real n * \
|
|
775 |
\ (real (n - Suc 0) * abs (c n) * inverse r) * r ^ (n - Suc 0))")]
|
|
776 |
(CLAIM "(a = b) ==> a ==> b") 1 THEN assume_tac 2);
|
|
777 |
by (res_inst_tac [("f","summable")] arg_cong 1 THEN rtac ext 1);
|
|
778 |
(* 75 *)
|
|
779 |
by (case_tac "n" 1);
|
|
780 |
by Auto_tac;
|
|
781 |
by (case_tac "nat" 1);
|
|
782 |
by Auto_tac;
|
|
783 |
by (dtac (abs_ge_zero RS order_le_less_trans) 1);
|
|
784 |
by (auto_tac (claset(),simpset() addsimps [CLAIM_SIMP
|
|
785 |
"(a::real) * (b * (c * d)) = a * (b * c) * d"
|
|
786 |
real_mult_ac]));
|
|
787 |
by (dtac (abs_ge_zero RS order_le_less_trans) 1);
|
|
788 |
by (asm_full_simp_tac (simpset() addsimps [abs_mult,real_mult_assoc]) 1);
|
|
789 |
by (rtac real_mult_le_le_mono1 1);
|
|
790 |
by (rtac (real_add_commute RS subst) 2);
|
|
791 |
by (simp_tac (simpset() addsimps [real_mult_assoc RS sym]) 2);
|
|
792 |
by (rtac lemma_termdiff3 2);
|
|
793 |
by (auto_tac (claset() addIs [(abs_triangle_ineq RS real_le_trans)],
|
|
794 |
simpset()));
|
|
795 |
by (arith_tac 1);
|
|
796 |
qed "termdiffs";
|
|
797 |
|
|
798 |
(* ------------------------------------------------------------------------ *)
|
|
799 |
(* Formal derivatives of exp, sin, and cos series *)
|
|
800 |
(* ------------------------------------------------------------------------ *)
|
|
801 |
|
|
802 |
Goalw [diffs_def]
|
|
803 |
"diffs (%n. inverse(real (fact n))) = (%n. inverse(real (fact n)))";
|
|
804 |
by (rtac ext 1);
|
12486
|
805 |
by (stac fact_Suc 1);
|
|
806 |
by (stac real_of_nat_mult 1);
|
|
807 |
by (stac real_inverse_distrib 1);
|
12196
|
808 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym]));
|
|
809 |
qed "exp_fdiffs";
|
|
810 |
|
|
811 |
Goalw [diffs_def,real_divide_def]
|
|
812 |
"diffs(%n. if even n then 0 \
|
|
813 |
\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) \
|
|
814 |
\ = (%n. if even n then \
|
|
815 |
\ (- 1) ^ (n div 2)/(real (fact n)) \
|
|
816 |
\ else 0)";
|
|
817 |
by (rtac ext 1);
|
12486
|
818 |
by (stac fact_Suc 1);
|
|
819 |
by (stac real_of_nat_mult 1);
|
|
820 |
by (stac even_Suc 1);
|
|
821 |
by (stac real_inverse_distrib 1);
|
12196
|
822 |
by Auto_tac;
|
|
823 |
qed "sin_fdiffs";
|
|
824 |
|
|
825 |
Goalw [diffs_def,real_divide_def]
|
|
826 |
"diffs(%n. if even n then 0 \
|
|
827 |
\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) n \
|
|
828 |
\ = (if even n then \
|
|
829 |
\ (- 1) ^ (n div 2)/(real (fact n)) \
|
|
830 |
\ else 0)";
|
12486
|
831 |
by (stac fact_Suc 1);
|
|
832 |
by (stac real_of_nat_mult 1);
|
|
833 |
by (stac even_Suc 1);
|
|
834 |
by (stac real_inverse_distrib 1);
|
12196
|
835 |
by Auto_tac;
|
|
836 |
qed "sin_fdiffs2";
|
|
837 |
|
|
838 |
(* thms in EvenOdd needed *)
|
|
839 |
Goalw [diffs_def,real_divide_def]
|
|
840 |
"diffs(%n. if even n then \
|
|
841 |
\ (- 1) ^ (n div 2)/(real (fact n)) else 0) \
|
|
842 |
\ = (%n. - (if even n then 0 \
|
|
843 |
\ else (- 1) ^ ((n - Suc 0)div 2)/(real (fact n))))";
|
|
844 |
by (rtac ext 1);
|
12486
|
845 |
by (stac fact_Suc 1);
|
|
846 |
by (stac real_of_nat_mult 1);
|
|
847 |
by (stac even_Suc 1);
|
|
848 |
by (stac real_inverse_distrib 1);
|
12196
|
849 |
by (res_inst_tac [("z1","real (Suc n)")] (real_mult_commute RS ssubst) 1);
|
|
850 |
by (res_inst_tac [("z1","inverse(real (Suc n))")]
|
|
851 |
(real_mult_commute RS ssubst) 1);
|
|
852 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
|
|
853 |
odd_not_even RS sym,odd_Suc_mult_two_ex]));
|
|
854 |
qed "cos_fdiffs";
|
|
855 |
|
|
856 |
|
|
857 |
Goalw [diffs_def,real_divide_def]
|
|
858 |
"diffs(%n. if even n then \
|
|
859 |
\ (- 1) ^ (n div 2)/(real (fact n)) else 0) n\
|
|
860 |
\ = - (if even n then 0 \
|
|
861 |
\ else (- 1) ^ ((n - Suc 0)div 2)/(real (fact n)))";
|
12486
|
862 |
by (stac fact_Suc 1);
|
|
863 |
by (stac real_of_nat_mult 1);
|
|
864 |
by (stac even_Suc 1);
|
|
865 |
by (stac real_inverse_distrib 1);
|
12196
|
866 |
by (res_inst_tac [("z1","real (Suc n)")] (real_mult_commute RS ssubst) 1);
|
|
867 |
by (res_inst_tac [("z1","inverse (real (Suc n))")]
|
|
868 |
(real_mult_commute RS ssubst) 1);
|
|
869 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,
|
|
870 |
odd_not_even RS sym,odd_Suc_mult_two_ex]));
|
|
871 |
qed "cos_fdiffs2";
|
|
872 |
|
|
873 |
(* ------------------------------------------------------------------------ *)
|
|
874 |
(* Now at last we can get the derivatives of exp, sin and cos *)
|
|
875 |
(* ------------------------------------------------------------------------ *)
|
|
876 |
|
|
877 |
Goal "- sin x = suminf(%n. - ((if even n then 0 \
|
|
878 |
\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * x ^ n))";
|
|
879 |
by (auto_tac (claset() addSIs [sums_unique,sums_minus,sin_converges],
|
|
880 |
simpset()));
|
|
881 |
qed "lemma_sin_minus";
|
|
882 |
|
|
883 |
Goal "exp = (%x. suminf (%n. inverse (real (fact n)) * x ^ n))";
|
|
884 |
by (auto_tac (claset() addSIs [ext],simpset() addsimps [exp_def]));
|
|
885 |
val lemma_exp_ext = result();
|
|
886 |
|
|
887 |
Goalw [exp_def] "DERIV exp x :> exp(x)";
|
12486
|
888 |
by (stac lemma_exp_ext 1);
|
12196
|
889 |
by (subgoal_tac "DERIV (%u. suminf (%n. inverse (real (fact n)) * u ^ n)) x \
|
|
890 |
\ :> suminf (%n. diffs (%n. inverse (real (fact n))) n * x ^ n)" 1);
|
|
891 |
by (res_inst_tac [("K","1 + abs(x)")] termdiffs 2);
|
|
892 |
by (auto_tac (claset() addIs [exp_converges RS sums_summable],
|
|
893 |
simpset() addsimps [exp_fdiffs]));
|
|
894 |
by (arith_tac 1);
|
|
895 |
qed "DERIV_exp";
|
|
896 |
Addsimps [DERIV_exp];
|
|
897 |
|
|
898 |
Goal "sin = (%x. suminf(%n. (if even n then 0 \
|
|
899 |
\ else (- 1) ^ ((n - Suc 0) div 2)/(real (fact n))) * \
|
|
900 |
\ x ^ n))";
|
|
901 |
by (auto_tac (claset() addSIs [ext],simpset() addsimps [sin_def]));
|
|
902 |
val lemma_sin_ext = result();
|
|
903 |
|
|
904 |
Goal "cos = (%x. suminf(%n. (if even n then \
|
|
905 |
\ (- 1) ^ (n div 2)/(real (fact n)) \
|
|
906 |
\ else 0) * x ^ n))";
|
|
907 |
by (auto_tac (claset() addSIs [ext],simpset() addsimps [cos_def]));
|
|
908 |
val lemma_cos_ext = result();
|
|
909 |
|
|
910 |
Goalw [cos_def] "DERIV sin x :> cos(x)";
|
12486
|
911 |
by (stac lemma_sin_ext 1);
|
12196
|
912 |
by (auto_tac (claset(),simpset() addsimps [sin_fdiffs2 RS sym]));
|
|
913 |
by (res_inst_tac [("K","1 + abs(x)")] termdiffs 1);
|
|
914 |
by (auto_tac (claset() addIs [sin_converges, cos_converges, sums_summable]
|
|
915 |
addSIs [sums_minus RS sums_summable],
|
|
916 |
simpset() addsimps [cos_fdiffs,sin_fdiffs]));
|
|
917 |
by (arith_tac 1);
|
|
918 |
qed "DERIV_sin";
|
|
919 |
Addsimps [DERIV_sin];
|
|
920 |
|
|
921 |
Goal "DERIV cos x :> -sin(x)";
|
12486
|
922 |
by (stac lemma_cos_ext 1);
|
12196
|
923 |
by (auto_tac (claset(),simpset() addsimps [lemma_sin_minus,
|
|
924 |
cos_fdiffs2 RS sym,real_minus_mult_eq1]));
|
|
925 |
by (res_inst_tac [("K","1 + abs(x)")] termdiffs 1);
|
|
926 |
by (auto_tac (claset() addIs [sin_converges,cos_converges, sums_summable]
|
|
927 |
addSIs [sums_minus RS sums_summable],
|
|
928 |
simpset() addsimps [cos_fdiffs,sin_fdiffs,diffs_minus]));
|
|
929 |
by (arith_tac 1);
|
|
930 |
qed "DERIV_cos";
|
|
931 |
Addsimps [DERIV_cos];
|
|
932 |
|
|
933 |
(* ------------------------------------------------------------------------ *)
|
|
934 |
(* Properties of the exponential function *)
|
|
935 |
(* ------------------------------------------------------------------------ *)
|
|
936 |
|
|
937 |
Goalw [exp_def] "exp 0 = 1";
|
|
938 |
by (rtac (CLAIM_SIMP "sumr 0 1 (%n. inverse (real (fact n)) * 0 ^ n) = 1"
|
|
939 |
[real_of_nat_one] RS subst) 1);
|
|
940 |
by (rtac ((series_zero RS sums_unique) RS sym) 1);
|
|
941 |
by (Step_tac 1);
|
|
942 |
by (case_tac "m" 1);
|
|
943 |
by (Auto_tac);
|
|
944 |
qed "exp_zero";
|
|
945 |
Addsimps [exp_zero];
|
|
946 |
|
|
947 |
Goal "0 <= x ==> (1 + x) <= exp(x)";
|
|
948 |
by (dtac real_le_imp_less_or_eq 1 THEN Auto_tac);
|
|
949 |
by (rewtac exp_def);
|
|
950 |
by (rtac real_le_trans 1);
|
|
951 |
by (res_inst_tac [("n","2"),("f","(%n. inverse (real (fact n)) * x ^ n)")]
|
|
952 |
series_pos_le 2);
|
|
953 |
by (auto_tac (claset() addIs [summable_exp],simpset()
|
|
954 |
addsimps [numeral_2_eq_2,realpow_ge_zero,real_0_le_mult_iff]));
|
|
955 |
qed "exp_ge_add_one_self";
|
|
956 |
Addsimps [exp_ge_add_one_self];
|
|
957 |
|
|
958 |
Goal "0 < x ==> 1 < exp x";
|
|
959 |
by (rtac order_less_le_trans 1);
|
|
960 |
by (rtac exp_ge_add_one_self 2);
|
|
961 |
by (Auto_tac);
|
|
962 |
qed "exp_gt_one";
|
|
963 |
Addsimps [exp_gt_one];
|
|
964 |
|
|
965 |
Goal "DERIV (%x. exp (x + y)) x :> exp(x + y)";
|
|
966 |
by (auto_tac (claset(),simpset() addsimps
|
|
967 |
[CLAIM_SIMP "(%x. exp (x + y)) = exp o (%x. x + y)" [ext]]));
|
|
968 |
by (rtac (real_mult_1_right RS subst) 1);
|
|
969 |
by (rtac DERIV_chain 1);
|
|
970 |
by (rtac (real_add_zero_right RS subst) 2);
|
|
971 |
by (rtac DERIV_add 2);
|
|
972 |
by Auto_tac;
|
|
973 |
qed "DERIV_exp_add_const";
|
|
974 |
Addsimps [DERIV_exp_add_const];
|
|
975 |
|
|
976 |
Goal "DERIV (%x. exp (-x)) x :> - exp(-x)";
|
|
977 |
by (auto_tac (claset(),simpset() addsimps
|
|
978 |
[CLAIM_SIMP "(%x. exp(-x)) = exp o (%x. - x)" [ext]]));
|
|
979 |
by (rtac (real_mult_1_right RS subst) 1);
|
|
980 |
by (rtac (real_minus_mult_eq1 RS subst) 1);
|
12486
|
981 |
by (stac real_minus_mult_eq2 1);
|
12196
|
982 |
by (rtac DERIV_chain 1);
|
|
983 |
by (rtac DERIV_minus 2);
|
|
984 |
by Auto_tac;
|
|
985 |
qed "DERIV_exp_minus";
|
|
986 |
Addsimps [DERIV_exp_minus];
|
|
987 |
|
|
988 |
Goal "DERIV (%x. exp (x + y) * exp (- x)) x :> 0";
|
|
989 |
by (cut_inst_tac [("x","x"),("y2","y")] ([DERIV_exp_add_const,
|
|
990 |
DERIV_exp_minus] MRS DERIV_mult) 1);
|
|
991 |
by (auto_tac (claset(),simpset() addsimps real_mult_ac));
|
|
992 |
qed "DERIV_exp_exp_zero";
|
|
993 |
Addsimps [DERIV_exp_exp_zero];
|
|
994 |
|
|
995 |
Goal "exp(x + y)*exp(-x) = exp(y)";
|
|
996 |
by (cut_inst_tac [("x","x"),("y2","y"),("y","0")]
|
|
997 |
((CLAIM "ALL x. DERIV (%x. exp (x + y) * exp (- x)) x :> 0") RS
|
|
998 |
DERIV_isconst_all) 1);
|
|
999 |
by (Auto_tac);
|
|
1000 |
qed "exp_add_mult_minus";
|
|
1001 |
Addsimps [exp_add_mult_minus];
|
|
1002 |
|
|
1003 |
Goal "exp(x)*exp(-x) = 1";
|
|
1004 |
by (cut_inst_tac [("x","x"),("y","0")] exp_add_mult_minus 1);
|
|
1005 |
by (Asm_full_simp_tac 1);
|
|
1006 |
qed "exp_mult_minus";
|
|
1007 |
Addsimps [exp_mult_minus];
|
|
1008 |
|
|
1009 |
Goal "exp(-x)*exp(x) = 1";
|
|
1010 |
by (simp_tac (simpset() addsimps [real_mult_commute]) 1);
|
|
1011 |
qed "exp_mult_minus2";
|
|
1012 |
Addsimps [exp_mult_minus2];
|
|
1013 |
|
|
1014 |
Goal "exp(-x) = inverse(exp(x))";
|
|
1015 |
by (auto_tac (claset() addIs [real_inverse_unique],simpset()));
|
|
1016 |
qed "exp_minus";
|
|
1017 |
|
|
1018 |
Goal "exp(x + y) = exp(x) * exp(y)";
|
|
1019 |
by (cut_inst_tac [("x1","x"),("y1","y"),("z","exp x")]
|
|
1020 |
(exp_add_mult_minus RS (CLAIM "x = y ==> z * y = z * (x::real)")) 1);
|
|
1021 |
by (asm_full_simp_tac (simpset() delsimps [exp_add_mult_minus]
|
|
1022 |
addsimps real_mult_ac) 1);
|
|
1023 |
qed "exp_add";
|
|
1024 |
|
|
1025 |
Goal "0 <= exp x";
|
|
1026 |
by (res_inst_tac [("t","x")] (real_sum_of_halves RS subst) 1);
|
12486
|
1027 |
by (stac exp_add 1 THEN Auto_tac);
|
12196
|
1028 |
qed "exp_ge_zero";
|
|
1029 |
Addsimps [exp_ge_zero];
|
|
1030 |
|
|
1031 |
Goal "exp x ~= 0";
|
|
1032 |
by (cut_inst_tac [("x","x")] exp_mult_minus2 1);
|
|
1033 |
by (auto_tac (claset(),simpset() delsimps [exp_mult_minus2]));
|
|
1034 |
qed "exp_not_eq_zero";
|
|
1035 |
Addsimps [exp_not_eq_zero];
|
|
1036 |
|
|
1037 |
Goal "0 < exp x";
|
|
1038 |
by (simp_tac (simpset() addsimps
|
|
1039 |
[CLAIM_SIMP "(x < y) = (x <= y & y ~= (x::real))" [real_le_less]]) 1);
|
|
1040 |
qed "exp_gt_zero";
|
|
1041 |
Addsimps [exp_gt_zero];
|
|
1042 |
|
|
1043 |
Goal "0 < inverse(exp x)";
|
|
1044 |
by (auto_tac (claset() addIs [real_inverse_gt_0],simpset()));
|
|
1045 |
qed "inv_exp_gt_zero";
|
|
1046 |
Addsimps [inv_exp_gt_zero];
|
|
1047 |
|
|
1048 |
Goal "abs(exp x) = exp x";
|
|
1049 |
by (auto_tac (claset(),simpset() addsimps [abs_eqI2]));
|
|
1050 |
qed "abs_exp_cancel";
|
|
1051 |
Addsimps [abs_exp_cancel];
|
|
1052 |
|
|
1053 |
Goal "exp(real n * x) = exp(x) ^ n";
|
|
1054 |
by (induct_tac "n" 1);
|
|
1055 |
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc,
|
|
1056 |
real_add_mult_distrib2,exp_add,real_mult_commute]));
|
|
1057 |
qed "exp_real_of_nat_mult";
|
|
1058 |
|
|
1059 |
Goalw [real_diff_def,real_divide_def]
|
|
1060 |
"exp(x - y) = exp(x)/(exp y)";
|
|
1061 |
by (simp_tac (simpset() addsimps [exp_add,exp_minus]) 1);
|
|
1062 |
qed "exp_diff";
|
|
1063 |
|
|
1064 |
Goal "x < y ==> exp x < exp y";
|
|
1065 |
by (dtac ((real_less_sum_gt_zero) RS exp_gt_one) 1);
|
|
1066 |
by (multr_by_tac "inverse(exp x)" 1);
|
|
1067 |
by (auto_tac (claset(),simpset() addsimps [exp_add,exp_minus]));
|
|
1068 |
qed "exp_less_mono";
|
|
1069 |
|
|
1070 |
Goal "exp x < exp y ==> x < y";
|
|
1071 |
by (EVERY1[rtac ccontr, dtac real_leI, dtac real_le_imp_less_or_eq]);
|
|
1072 |
by (auto_tac (claset() addDs [exp_less_mono],simpset()));
|
|
1073 |
qed "exp_less_cancel";
|
|
1074 |
|
|
1075 |
Goal "(exp(x) < exp(y)) = (x < y)";
|
|
1076 |
by (auto_tac (claset() addIs [exp_less_mono,exp_less_cancel],simpset()));
|
|
1077 |
qed "exp_less_cancel_iff";
|
|
1078 |
AddIffs [exp_less_cancel_iff];
|
|
1079 |
|
|
1080 |
Goalw [real_le_def] "(exp(x) <= exp(y)) = (x <= y)";
|
|
1081 |
by (Auto_tac);
|
|
1082 |
qed "exp_le_cancel_iff";
|
|
1083 |
AddIffs [exp_le_cancel_iff];
|
|
1084 |
|
|
1085 |
Goal "(exp x = exp y) = (x = y)";
|
|
1086 |
by (auto_tac (claset(),simpset() addsimps
|
|
1087 |
[CLAIM "(x = (y::real)) = (x <= y & y <= x)"]));
|
|
1088 |
qed "exp_inj_iff";
|
|
1089 |
AddIffs [exp_inj_iff];
|
|
1090 |
|
|
1091 |
Goal "1 <= y ==> EX x. 0 <= x & x <= y - 1 & exp(x) = y";
|
|
1092 |
by (rtac IVT 1);
|
|
1093 |
by (auto_tac (claset() addIs [DERIV_exp RS DERIV_isCont],simpset()
|
|
1094 |
addsimps [real_le_diff_eq]));
|
|
1095 |
by (dtac (CLAIM_SIMP "x <= y ==> (0::real) <= y - x" [real_le_diff_eq]) 1);
|
|
1096 |
by (dtac exp_ge_add_one_self 1);
|
|
1097 |
by (Asm_full_simp_tac 1);
|
|
1098 |
qed "lemma_exp_total";
|
|
1099 |
|
|
1100 |
Goal "0 < y ==> EX x. exp x = y";
|
|
1101 |
by (res_inst_tac [("R1.0","1"),("R2.0","y")] real_linear_less2 1);
|
|
1102 |
by (dtac (order_less_imp_le RS lemma_exp_total) 1);
|
|
1103 |
by (res_inst_tac [("x","0")] exI 2);
|
|
1104 |
by (ftac real_inverse_gt_one 3);
|
|
1105 |
by (dtac (order_less_imp_le RS lemma_exp_total) 4);
|
|
1106 |
by (Step_tac 3);
|
|
1107 |
by (res_inst_tac [("x","-x")] exI 3);
|
|
1108 |
by (auto_tac (claset(),simpset() addsimps [exp_minus]));
|
|
1109 |
qed "exp_total";
|
|
1110 |
|
|
1111 |
(* ------------------------------------------------------------------------ *)
|
|
1112 |
(* Properties of the logarithmic function *)
|
|
1113 |
(* ------------------------------------------------------------------------ *)
|
|
1114 |
|
|
1115 |
Goalw [ln_def] "ln(exp x) = x";
|
|
1116 |
by (Simp_tac 1);
|
|
1117 |
qed "ln_exp";
|
|
1118 |
Addsimps [ln_exp];
|
|
1119 |
|
|
1120 |
Goal "(exp(ln x) = x) = (0 < x)";
|
|
1121 |
by (auto_tac (claset() addDs [exp_total],simpset()));
|
|
1122 |
by (dtac subst 1);
|
|
1123 |
by (Auto_tac);
|
|
1124 |
qed "exp_ln_iff";
|
|
1125 |
Addsimps [exp_ln_iff];
|
|
1126 |
|
|
1127 |
Goal "[| 0 < x; 0 < y |] ==> ln(x * y) = ln(x) + ln(y)";
|
|
1128 |
by (rtac (exp_inj_iff RS iffD1) 1);
|
|
1129 |
by (ftac (real_mult_order) 1);
|
|
1130 |
by (auto_tac (claset(),simpset() addsimps [exp_add,exp_ln_iff RS sym]
|
|
1131 |
delsimps [exp_inj_iff,exp_ln_iff]));
|
|
1132 |
qed "ln_mult";
|
|
1133 |
|
|
1134 |
Goal "[| 0 < x; 0 < y |] ==> (ln x = ln y) = (x = y)";
|
13601
|
1135 |
by (auto_tac (claset() addSDs [(exp_ln_iff RS iffD2 RS sym)],simpset()));
|
12196
|
1136 |
qed "ln_inj_iff";
|
|
1137 |
Addsimps [ln_inj_iff];
|
|
1138 |
|
|
1139 |
Goal "ln 1 = 0";
|
|
1140 |
by (rtac (exp_inj_iff RS iffD1) 1);
|
|
1141 |
by Auto_tac;
|
|
1142 |
qed "ln_one";
|
|
1143 |
Addsimps [ln_one];
|
|
1144 |
|
|
1145 |
Goal "0 < x ==> ln(inverse x) = - ln x";
|
|
1146 |
by (res_inst_tac [("x1","ln x")] (real_add_left_cancel RS iffD1) 1);
|
|
1147 |
by (auto_tac (claset(),simpset() addsimps [real_inverse_gt_0,ln_mult RS sym]));
|
|
1148 |
qed "ln_inverse";
|
|
1149 |
|
|
1150 |
Goalw [real_divide_def]
|
|
1151 |
"[|0 < x; 0 < y|] ==> ln(x/y) = ln x - ln y";
|
|
1152 |
by (auto_tac (claset(),simpset() addsimps [real_inverse_gt_0,
|
|
1153 |
ln_mult,ln_inverse]));
|
|
1154 |
qed "ln_div";
|
|
1155 |
|
|
1156 |
Goal "[| 0 < x; 0 < y|] ==> (ln x < ln y) = (x < y)";
|
|
1157 |
by (REPEAT(dtac (exp_ln_iff RS iffD2) 1));
|
|
1158 |
by (REPEAT(dtac subst 1 THEN assume_tac 2));
|
|
1159 |
by (Simp_tac 1);
|
|
1160 |
qed "ln_less_cancel_iff";
|
|
1161 |
Addsimps [ln_less_cancel_iff];
|
|
1162 |
|
|
1163 |
Goalw [real_le_def] "[| 0 < x; 0 < y|] ==> (ln x <= ln y) = (x <= y)";
|
|
1164 |
by (Auto_tac);
|
|
1165 |
qed "ln_le_cancel_iff";
|
|
1166 |
Addsimps [ln_le_cancel_iff];
|
|
1167 |
|
|
1168 |
Goal "0 < x ==> ln(x ^ n) = real n * ln(x)";
|
|
1169 |
by (auto_tac (claset() addSDs [exp_total],simpset()
|
|
1170 |
addsimps [exp_real_of_nat_mult RS sym]));
|
|
1171 |
qed "ln_realpow";
|
|
1172 |
|
|
1173 |
Goal "0 <= x ==> ln(1 + x) <= x";
|
|
1174 |
by (rtac (ln_exp RS subst) 1);
|
|
1175 |
by (rtac (ln_le_cancel_iff RS iffD2) 1);
|
|
1176 |
by Auto_tac;
|
|
1177 |
qed "ln_add_one_self_le_self";
|
|
1178 |
Addsimps [ln_add_one_self_le_self];
|
|
1179 |
|
|
1180 |
Goal "0 < x ==> ln x < x";
|
|
1181 |
by (rtac order_less_le_trans 1);
|
|
1182 |
by (rtac ln_add_one_self_le_self 2);
|
|
1183 |
by (rtac (ln_less_cancel_iff RS iffD2) 1);
|
|
1184 |
by Auto_tac;
|
|
1185 |
qed "ln_less_self";
|
|
1186 |
Addsimps [ln_less_self];
|
|
1187 |
|
|
1188 |
Goal "1 <= x ==> 0 <= ln x";
|
|
1189 |
by (subgoal_tac "0 < x" 1);
|
|
1190 |
by (rtac order_less_le_trans 2 THEN assume_tac 3);
|
|
1191 |
by (rtac (exp_le_cancel_iff RS iffD1) 1);
|
|
1192 |
by (auto_tac (claset(),simpset() addsimps
|
|
1193 |
[exp_ln_iff RS sym] delsimps [exp_ln_iff]));
|
|
1194 |
qed "ln_ge_zero";
|
|
1195 |
Addsimps [ln_ge_zero];
|
|
1196 |
|
|
1197 |
Goal "1 < x ==> 0 < ln x";
|
|
1198 |
by (rtac (exp_less_cancel_iff RS iffD1) 1);
|
|
1199 |
by (rtac (exp_ln_iff RS iffD2 RS ssubst) 1);
|
|
1200 |
by Auto_tac;
|
|
1201 |
qed "ln_gt_zero";
|
|
1202 |
Addsimps [ln_gt_zero];
|
|
1203 |
|
|
1204 |
Goal "[| 0 < x; x ~= 1 |] ==> ln x ~= 0";
|
|
1205 |
by (Step_tac 1);
|
|
1206 |
by (dtac (exp_inj_iff RS iffD2) 1);
|
|
1207 |
by (dtac (exp_ln_iff RS iffD2) 1);
|
|
1208 |
by Auto_tac;
|
|
1209 |
qed "ln_not_eq_zero";
|
|
1210 |
Addsimps [ln_not_eq_zero];
|
|
1211 |
|
|
1212 |
Goal "[| 0 < x; x < 1 |] ==> ln x < 0";
|
|
1213 |
by (rtac (exp_less_cancel_iff RS iffD1) 1);
|
|
1214 |
by (auto_tac (claset(),simpset() addsimps [exp_ln_iff RS sym]
|
|
1215 |
delsimps [exp_ln_iff]));
|
|
1216 |
qed "ln_less_zero";
|
|
1217 |
|
|
1218 |
Goal "exp u = x ==> ln x = u";
|
|
1219 |
by Auto_tac;
|
|
1220 |
qed "exp_ln_eq";
|
|
1221 |
|
|
1222 |
Addsimps [hypreal_less_not_refl];
|
|
1223 |
|
|
1224 |
(* ------------------------------------------------------------------------ *)
|
|
1225 |
(* Basic properties of the trig functions *)
|
|
1226 |
(* ------------------------------------------------------------------------ *)
|
|
1227 |
|
|
1228 |
Goalw [sin_def] "sin 0 = 0";
|
|
1229 |
by (auto_tac (claset() addSIs [(sums_unique RS sym),
|
|
1230 |
LIMSEQ_const],simpset() addsimps [sums_def]));
|
|
1231 |
qed "sin_zero";
|
|
1232 |
Addsimps [sin_zero];
|
|
1233 |
|
|
1234 |
Goal "(ALL m. n <= m --> f m = 0) --> f sums sumr 0 n f";
|
|
1235 |
by (auto_tac (claset() addIs [series_zero],simpset()));
|
|
1236 |
qed "lemma_series_zero2";
|
|
1237 |
|
|
1238 |
Goalw [cos_def] "cos 0 = 1";
|
|
1239 |
by (rtac (sums_unique RS sym) 1);
|
|
1240 |
by (cut_inst_tac [("n","1"),("f","(%n. (if even n then (- 1) ^ (n div 2)/ \
|
|
1241 |
\ (real (fact n)) else 0) * 0 ^ n)")] lemma_series_zero2 1);
|
|
1242 |
by Auto_tac;
|
|
1243 |
qed "cos_zero";
|
|
1244 |
Addsimps [cos_zero];
|
|
1245 |
|
|
1246 |
Goal "DERIV (%x. sin(x)*sin(x)) x :> cos(x) * sin(x) + cos(x) * sin(x)";
|
|
1247 |
by (rtac DERIV_mult 1 THEN Auto_tac);
|
|
1248 |
qed "DERIV_sin_sin_mult";
|
|
1249 |
Addsimps [DERIV_sin_sin_mult];
|
|
1250 |
|
|
1251 |
Goal "DERIV (%x. sin(x)*sin(x)) x :> 2 * cos(x) * sin(x)";
|
|
1252 |
by (cut_inst_tac [("x","x")] DERIV_sin_sin_mult 1);
|
|
1253 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc]));
|
|
1254 |
qed "DERIV_sin_sin_mult2";
|
|
1255 |
Addsimps [DERIV_sin_sin_mult2];
|
|
1256 |
|
|
1257 |
Goal "DERIV (%x. sin(x) ^ 2) x :> cos(x) * sin(x) + cos(x) * sin(x)";
|
|
1258 |
by (auto_tac (claset(),
|
|
1259 |
simpset() addsimps [numeral_2_eq_2, real_mult_assoc RS sym]));
|
|
1260 |
qed "DERIV_sin_realpow2";
|
|
1261 |
Addsimps [DERIV_sin_realpow2];
|
|
1262 |
|
|
1263 |
Goal "DERIV (%x. sin(x) ^ 2) x :> 2 * cos(x) * sin(x)";
|
|
1264 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
1265 |
qed "DERIV_sin_realpow2a";
|
|
1266 |
Addsimps [ DERIV_sin_realpow2a];
|
|
1267 |
|
|
1268 |
Goal "DERIV (%x. cos(x)*cos(x)) x :> -sin(x) * cos(x) + -sin(x) * cos(x)";
|
|
1269 |
by (rtac DERIV_mult 1 THEN Auto_tac);
|
|
1270 |
qed "DERIV_cos_cos_mult";
|
|
1271 |
Addsimps [DERIV_cos_cos_mult];
|
|
1272 |
|
|
1273 |
Goal "DERIV (%x. cos(x)*cos(x)) x :> -2 * cos(x) * sin(x)";
|
|
1274 |
by (cut_inst_tac [("x","x")] DERIV_cos_cos_mult 1);
|
|
1275 |
by (auto_tac (claset(),simpset() addsimps real_mult_ac));
|
|
1276 |
qed "DERIV_cos_cos_mult2";
|
|
1277 |
Addsimps [DERIV_cos_cos_mult2];
|
|
1278 |
|
|
1279 |
Goal "DERIV (%x. cos(x) ^ 2) x :> -sin(x) * cos(x) + -sin(x) * cos(x)";
|
|
1280 |
by (auto_tac (claset(),
|
|
1281 |
simpset() addsimps [numeral_2_eq_2, real_mult_assoc RS sym]));
|
|
1282 |
qed "DERIV_cos_realpow2";
|
|
1283 |
Addsimps [DERIV_cos_realpow2];
|
|
1284 |
|
|
1285 |
Goal "DERIV (%x. cos(x) ^ 2) x :> -2 * cos(x) * sin(x)";
|
|
1286 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
1287 |
qed "DERIV_cos_realpow2a";
|
|
1288 |
Addsimps [DERIV_cos_realpow2a];
|
|
1289 |
|
|
1290 |
Goal "[| DERIV f x :> D; D = E |] ==> DERIV f x :> E";
|
|
1291 |
by (Auto_tac);
|
|
1292 |
val lemma_DERIV_subst = result();
|
|
1293 |
|
|
1294 |
Goal "DERIV (%x. cos(x) ^ 2) x :> -(2 * cos(x) * sin(x))";
|
|
1295 |
by (rtac lemma_DERIV_subst 1);
|
|
1296 |
by (rtac DERIV_cos_realpow2a 1);
|
|
1297 |
by Auto_tac;
|
|
1298 |
qed "DERIV_cos_realpow2b";
|
|
1299 |
|
|
1300 |
(* most useful *)
|
|
1301 |
Goal "DERIV (%x. cos(x)*cos(x)) x :> -(2 * cos(x) * sin(x))";
|
|
1302 |
by (rtac lemma_DERIV_subst 1);
|
|
1303 |
by (rtac DERIV_cos_cos_mult2 1);
|
|
1304 |
by Auto_tac;
|
|
1305 |
qed "DERIV_cos_cos_mult3";
|
|
1306 |
Addsimps [DERIV_cos_cos_mult3];
|
|
1307 |
|
|
1308 |
Goalw [real_diff_def]
|
|
1309 |
"ALL x. DERIV (%x. sin(x) ^ 2 + cos(x) ^ 2) x :> \
|
|
1310 |
\ (2*cos(x)*sin(x) - 2*cos(x)*sin(x))";
|
|
1311 |
by (Step_tac 1);
|
|
1312 |
by (rtac DERIV_add 1);
|
|
1313 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
1314 |
qed "DERIV_sin_circle_all";
|
|
1315 |
|
|
1316 |
Goal "ALL x. DERIV (%x. sin(x) ^ 2 + cos(x) ^ 2) x :> 0";
|
|
1317 |
by (cut_facts_tac [DERIV_sin_circle_all] 1);
|
|
1318 |
by Auto_tac;
|
|
1319 |
qed "DERIV_sin_circle_all_zero";
|
|
1320 |
Addsimps [DERIV_sin_circle_all_zero];
|
|
1321 |
|
|
1322 |
Goal "(sin(x) ^ 2) + (cos(x) ^ 2) = 1";
|
|
1323 |
by (cut_inst_tac [("x","x"),("y","0")]
|
|
1324 |
(DERIV_sin_circle_all_zero RS DERIV_isconst_all) 1);
|
|
1325 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
1326 |
qed "sin_cos_squared_add";
|
|
1327 |
Addsimps [sin_cos_squared_add];
|
|
1328 |
|
|
1329 |
Goal "(cos(x) ^ 2) + (sin(x) ^ 2) = 1";
|
12486
|
1330 |
by (stac real_add_commute 1);
|
12196
|
1331 |
by (simp_tac (simpset() delsimps [realpow_Suc]) 1);
|
|
1332 |
qed "sin_cos_squared_add2";
|
|
1333 |
Addsimps [sin_cos_squared_add2];
|
|
1334 |
|
|
1335 |
Goal "cos x * cos x + sin x * sin x = 1";
|
|
1336 |
by (cut_inst_tac [("x","x")] sin_cos_squared_add2 1);
|
|
1337 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
1338 |
qed "sin_cos_squared_add3";
|
|
1339 |
Addsimps [sin_cos_squared_add3];
|
|
1340 |
|
|
1341 |
Goal "(sin(x) ^ 2) = 1 - (cos(x) ^ 2)";
|
|
1342 |
by (res_inst_tac [("x1","(cos(x) ^ 2)")] (real_add_right_cancel RS iffD1) 1);
|
|
1343 |
by (simp_tac (simpset() delsimps [realpow_Suc]) 1);
|
|
1344 |
qed "sin_squared_eq";
|
|
1345 |
|
|
1346 |
Goal "(cos(x) ^ 2) = 1 - (sin(x) ^ 2)";
|
|
1347 |
by (res_inst_tac [("x1","(sin(x) ^ 2)")] (real_add_right_cancel RS iffD1) 1);
|
|
1348 |
by (simp_tac (simpset() delsimps [realpow_Suc]) 1);
|
|
1349 |
qed "cos_squared_eq";
|
|
1350 |
|
|
1351 |
Goal "[| 1 < x; 0 <= y |] ==> 1 < x + (y::real)";
|
|
1352 |
by (arith_tac 1);
|
|
1353 |
qed "real_gt_one_ge_zero_add_less";
|
|
1354 |
|
|
1355 |
Goalw [real_le_def] "abs(sin x) <= 1";
|
|
1356 |
by (rtac notI 1);
|
|
1357 |
by (dtac realpow_two_gt_one 1);
|
|
1358 |
by (auto_tac (claset(),simpset() delsimps [realpow_Suc]));
|
|
1359 |
by (dres_inst_tac [("r1","cos x")] (realpow_two_le RSN
|
|
1360 |
(2, real_gt_one_ge_zero_add_less)) 1);
|
|
1361 |
by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym]
|
|
1362 |
delsimps [realpow_Suc]) 1);
|
|
1363 |
qed "abs_sin_le_one";
|
|
1364 |
Addsimps [abs_sin_le_one];
|
|
1365 |
|
|
1366 |
Goal "- 1 <= sin x";
|
|
1367 |
by (full_simp_tac (simpset() addsimps [simplify (simpset()) (abs_sin_le_one RS
|
|
1368 |
(abs_le_interval_iff RS iffD1))]) 1);
|
|
1369 |
qed "sin_ge_minus_one";
|
|
1370 |
Addsimps [sin_ge_minus_one];
|
|
1371 |
|
|
1372 |
Goal "-1 <= sin x";
|
|
1373 |
by (rtac (simplify (simpset()) sin_ge_minus_one) 1);
|
|
1374 |
qed "sin_ge_minus_one2";
|
|
1375 |
Addsimps [sin_ge_minus_one2];
|
|
1376 |
|
|
1377 |
Goal "sin x <= 1";
|
|
1378 |
by (full_simp_tac (simpset() addsimps [abs_sin_le_one RS
|
|
1379 |
(abs_le_interval_iff RS iffD1)]) 1);
|
|
1380 |
qed "sin_le_one";
|
|
1381 |
Addsimps [sin_le_one];
|
|
1382 |
|
|
1383 |
Goalw [real_le_def] "abs(cos x) <= 1";
|
|
1384 |
by (rtac notI 1);
|
|
1385 |
by (dtac realpow_two_gt_one 1);
|
|
1386 |
by (auto_tac (claset(),simpset() delsimps [realpow_Suc]));
|
|
1387 |
by (dres_inst_tac [("r1","sin x")] (realpow_two_le RSN
|
|
1388 |
(2, real_gt_one_ge_zero_add_less)) 1);
|
|
1389 |
by (asm_full_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym]
|
|
1390 |
delsimps [realpow_Suc]) 1);
|
|
1391 |
qed "abs_cos_le_one";
|
|
1392 |
Addsimps [abs_cos_le_one];
|
|
1393 |
|
|
1394 |
Goal "- 1 <= cos x";
|
|
1395 |
by (full_simp_tac (simpset() addsimps [simplify (simpset())(abs_cos_le_one RS
|
|
1396 |
(abs_le_interval_iff RS iffD1))]) 1);
|
|
1397 |
qed "cos_ge_minus_one";
|
|
1398 |
Addsimps [cos_ge_minus_one];
|
|
1399 |
|
|
1400 |
Goal "-1 <= cos x";
|
|
1401 |
by (rtac (simplify (simpset()) cos_ge_minus_one) 1);
|
|
1402 |
qed "cos_ge_minus_one2";
|
|
1403 |
Addsimps [cos_ge_minus_one2];
|
|
1404 |
|
|
1405 |
Goal "cos x <= 1";
|
|
1406 |
by (full_simp_tac (simpset() addsimps [abs_cos_le_one RS
|
|
1407 |
(abs_le_interval_iff RS iffD1)]) 1);
|
|
1408 |
qed "cos_le_one";
|
|
1409 |
Addsimps [cos_le_one];
|
|
1410 |
|
|
1411 |
Goal "DERIV g x :> m ==> \
|
|
1412 |
\ DERIV (%x. (g x) ^ n) x :> real n * (g x) ^ (n - 1) * m";
|
|
1413 |
by (rtac lemma_DERIV_subst 1);
|
|
1414 |
by (res_inst_tac [("f","(%x. x ^ n)")] DERIV_chain2 1);
|
|
1415 |
by (rtac DERIV_pow 1 THEN Auto_tac);
|
|
1416 |
qed "DERIV_fun_pow";
|
|
1417 |
|
|
1418 |
Goal "DERIV g x :> m ==> DERIV (%x. exp(g x)) x :> exp(g x) * m";
|
|
1419 |
by (rtac lemma_DERIV_subst 1);
|
|
1420 |
by (res_inst_tac [("f","exp")] DERIV_chain2 1);
|
|
1421 |
by (rtac DERIV_exp 1 THEN Auto_tac);
|
|
1422 |
qed "DERIV_fun_exp";
|
|
1423 |
|
|
1424 |
Goal "DERIV g x :> m ==> DERIV (%x. sin(g x)) x :> cos(g x) * m";
|
|
1425 |
by (rtac lemma_DERIV_subst 1);
|
|
1426 |
by (res_inst_tac [("f","sin")] DERIV_chain2 1);
|
|
1427 |
by (rtac DERIV_sin 1 THEN Auto_tac);
|
|
1428 |
qed "DERIV_fun_sin";
|
|
1429 |
|
|
1430 |
Goal "DERIV g x :> m ==> DERIV (%x. cos(g x)) x :> -sin(g x) * m";
|
|
1431 |
by (rtac lemma_DERIV_subst 1);
|
|
1432 |
by (res_inst_tac [("f","cos")] DERIV_chain2 1);
|
|
1433 |
by (rtac DERIV_cos 1 THEN Auto_tac);
|
|
1434 |
qed "DERIV_fun_cos";
|
|
1435 |
|
|
1436 |
(* FIXME: remove this quick, crude tactic *)
|
|
1437 |
exception DERIV_name;
|
|
1438 |
fun get_fun_name (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _)) = f
|
|
1439 |
| get_fun_name (_ $ (_ $ (Const ("Lim.deriv",_) $ Abs(_,_, Const (f,_) $ _) $ _ $ _))) = f
|
|
1440 |
| get_fun_name _ = raise DERIV_name;
|
|
1441 |
|
|
1442 |
val deriv_rulesI = [DERIV_Id,DERIV_const,DERIV_cos,DERIV_cmult,
|
|
1443 |
DERIV_sin, DERIV_exp, DERIV_inverse,DERIV_pow,
|
|
1444 |
DERIV_add, DERIV_diff, DERIV_mult, DERIV_minus,
|
|
1445 |
DERIV_inverse_fun,DERIV_quotient,DERIV_fun_pow,
|
|
1446 |
DERIV_fun_exp,DERIV_fun_sin,DERIV_fun_cos,
|
|
1447 |
DERIV_Id,DERIV_const,DERIV_cos];
|
|
1448 |
|
|
1449 |
|
|
1450 |
fun deriv_tac i = (resolve_tac deriv_rulesI i) ORELSE
|
|
1451 |
((rtac (read_instantiate [("f",get_fun_name (getgoal i))]
|
|
1452 |
DERIV_chain2) i) handle DERIV_name => no_tac);
|
|
1453 |
|
|
1454 |
val DERIV_tac = ALLGOALS(fn i => REPEAT(deriv_tac i));
|
|
1455 |
|
|
1456 |
(* lemma *)
|
|
1457 |
Goal "ALL x. \
|
|
1458 |
\ DERIV (%x. (sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + \
|
|
1459 |
\ (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2) x :> 0";
|
|
1460 |
by (Step_tac 1 THEN rtac lemma_DERIV_subst 1);
|
|
1461 |
by DERIV_tac;
|
|
1462 |
by (auto_tac (claset(),simpset() addsimps [real_diff_def,
|
|
1463 |
real_add_mult_distrib,real_add_mult_distrib2] @
|
|
1464 |
real_mult_ac @ real_add_ac));
|
|
1465 |
val lemma_DERIV_sin_cos_add = result();
|
|
1466 |
|
|
1467 |
Goal "(sin (x + y) - (sin x * cos y + cos x * sin y)) ^ 2 + \
|
|
1468 |
\ (cos (x + y) - (cos x * cos y - sin x * sin y)) ^ 2 = 0";
|
|
1469 |
by (cut_inst_tac [("y","0"),("x","x"),("y7","y")]
|
|
1470 |
(lemma_DERIV_sin_cos_add RS DERIV_isconst_all) 1);
|
|
1471 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
1472 |
qed "sin_cos_add";
|
|
1473 |
Addsimps [sin_cos_add];
|
|
1474 |
|
|
1475 |
Goal "sin (x + y) = sin x * cos y + cos x * sin y";
|
|
1476 |
by (cut_inst_tac [("x","x"),("y","y")] sin_cos_add 1);
|
|
1477 |
by (auto_tac (claset() addSDs [real_sum_squares_cancel_a],
|
|
1478 |
simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_add]));
|
|
1479 |
qed "sin_add";
|
|
1480 |
|
|
1481 |
Goal "cos (x + y) = cos x * cos y - sin x * sin y";
|
|
1482 |
by (cut_inst_tac [("x","x"),("y","y")] sin_cos_add 1);
|
|
1483 |
by (auto_tac (claset() addSDs [real_sum_squares_cancel_a],simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_add]));
|
|
1484 |
qed "cos_add";
|
|
1485 |
|
|
1486 |
Goal "ALL x. \
|
|
1487 |
\ DERIV (%x. (sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2) x :> 0";
|
|
1488 |
by (Step_tac 1 THEN rtac lemma_DERIV_subst 1);
|
|
1489 |
by DERIV_tac;
|
|
1490 |
by (auto_tac (claset(),simpset() addsimps [real_diff_def,
|
|
1491 |
real_add_mult_distrib,real_add_mult_distrib2]
|
|
1492 |
@ real_mult_ac @ real_add_ac));
|
|
1493 |
val lemma_DERIV_sin_cos_minus = result();
|
|
1494 |
|
|
1495 |
Goal "(sin(-x) + (sin x)) ^ 2 + (cos(-x) - (cos x)) ^ 2 = 0";
|
|
1496 |
by (cut_inst_tac [("y","0"),("x","x")]
|
|
1497 |
(lemma_DERIV_sin_cos_minus RS DERIV_isconst_all) 1);
|
|
1498 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
1499 |
qed "sin_cos_minus";
|
|
1500 |
Addsimps [sin_cos_minus];
|
|
1501 |
|
|
1502 |
Goal "sin (-x) = -sin(x)";
|
|
1503 |
by (cut_inst_tac [("x","x")] sin_cos_minus 1);
|
|
1504 |
by (auto_tac (claset() addSDs [real_sum_squares_cancel_a],
|
|
1505 |
simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_minus]));
|
|
1506 |
qed "sin_minus";
|
|
1507 |
Addsimps [sin_minus];
|
|
1508 |
|
|
1509 |
Goal "cos (-x) = cos(x)";
|
|
1510 |
by (cut_inst_tac [("x","x")] sin_cos_minus 1);
|
|
1511 |
by (auto_tac (claset() addSDs [real_sum_squares_cancel_a],
|
|
1512 |
simpset() addsimps [numeral_2_eq_2] delsimps [sin_cos_minus]));
|
|
1513 |
qed "cos_minus";
|
|
1514 |
Addsimps [cos_minus];
|
|
1515 |
|
|
1516 |
Goalw [real_diff_def] "sin (x - y) = sin x * cos y - cos x * sin y";
|
|
1517 |
by (simp_tac (simpset() addsimps [sin_add]) 1);
|
|
1518 |
qed "sin_diff";
|
|
1519 |
|
|
1520 |
Goal "sin (x - y) = cos y * sin x - sin y * cos x";
|
|
1521 |
by (simp_tac (simpset() addsimps [sin_diff,real_mult_commute]) 1);
|
|
1522 |
qed "sin_diff2";
|
|
1523 |
|
|
1524 |
Goalw [real_diff_def] "cos (x - y) = cos x * cos y + sin x * sin y";
|
|
1525 |
by (simp_tac (simpset() addsimps [cos_add]) 1);
|
|
1526 |
qed "cos_diff";
|
|
1527 |
|
|
1528 |
Goal "cos (x - y) = cos y * cos x + sin y * sin x";
|
|
1529 |
by (simp_tac (simpset() addsimps [cos_diff,real_mult_commute]) 1);
|
|
1530 |
qed "cos_diff2";
|
|
1531 |
|
|
1532 |
Goal "sin(2 * x) = 2* sin x * cos x";
|
|
1533 |
by (cut_inst_tac [("x","x"),("y","x")] sin_add 1);
|
|
1534 |
by Auto_tac;
|
|
1535 |
qed "sin_double";
|
|
1536 |
|
|
1537 |
Addsimps [sin_double];
|
|
1538 |
|
|
1539 |
Goal "cos(2* x) = (cos(x) ^ 2) - (sin(x) ^ 2)";
|
|
1540 |
by (cut_inst_tac [("x","x"),("y","x")] cos_add 1);
|
|
1541 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
1542 |
qed "cos_double";
|
|
1543 |
|
|
1544 |
(* ------------------------------------------------------------------------ *)
|
|
1545 |
(* Show that there's a least positive x with cos(x) = 0; hence define pi *)
|
|
1546 |
(* ------------------------------------------------------------------------ *)
|
|
1547 |
|
|
1548 |
Goal "(%n. (- 1) ^ n /(real (fact (2 * n + 1))) * \
|
|
1549 |
\ x ^ (2 * n + 1)) sums sin x";
|
|
1550 |
by (cut_inst_tac [("x2","x")] (CLAIM "0 < (2::nat)" RS ((sin_converges
|
|
1551 |
RS sums_summable) RS sums_group)) 1);
|
|
1552 |
by (auto_tac (claset(),simpset() addsimps mult_ac@[sin_def]));
|
|
1553 |
qed "sin_paired";
|
|
1554 |
|
|
1555 |
Goal "real (Suc (Suc (Suc (Suc 2)))) = \
|
|
1556 |
\ real (2::nat) * real (Suc 2)";
|
|
1557 |
by (simp_tac (simpset() addsimps [numeral_2_eq_2, real_of_nat_Suc]) 1);
|
|
1558 |
val lemma_real_of_nat_six_mult = result();
|
|
1559 |
|
|
1560 |
Goal "[|0 < x; x < 2 |] ==> 0 < sin x";
|
|
1561 |
by (cut_inst_tac [("x2","x")] (CLAIM "0 < (2::nat)" RS ((sin_paired
|
|
1562 |
RS sums_summable) RS sums_group)) 1);
|
|
1563 |
by (rotate_tac 2 1);
|
|
1564 |
by (dtac ((sin_paired RS sums_unique) RS ssubst) 1);
|
|
1565 |
by (auto_tac (claset(),simpset() delsimps [fact_Suc,realpow_Suc]));
|
|
1566 |
by (ftac sums_unique 1);
|
|
1567 |
by (auto_tac (claset(),simpset() delsimps [fact_Suc,realpow_Suc]));
|
|
1568 |
by (res_inst_tac [("n1","0")] (series_pos_less RSN (2,order_le_less_trans)) 1);
|
|
1569 |
by (auto_tac (claset(),simpset() delsimps [fact_Suc,realpow_Suc]));
|
|
1570 |
by (etac sums_summable 1);
|
|
1571 |
by (case_tac "m=0" 1);
|
|
1572 |
by (Asm_simp_tac 1);
|
|
1573 |
by (res_inst_tac [("z","real (Suc (Suc (Suc (Suc 2))))")]
|
|
1574 |
(CLAIM "[|(0::real)<z; z*x<z*y |] ==> x<y") 1);
|
|
1575 |
by (asm_simp_tac (simpset() addsimps [numeral_2_eq_2 RS sym, real_mult_assoc RS sym]) 2);
|
|
1576 |
by (stac (CLAIM "6 = 2 * (3::real)") 2);
|
|
1577 |
by (rtac real_mult_less_mono 2);
|
|
1578 |
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc] delsimps [fact_Suc]));
|
12486
|
1579 |
by (stac fact_Suc 1);
|
|
1580 |
by (stac fact_Suc 1);
|
|
1581 |
by (stac fact_Suc 1);
|
|
1582 |
by (stac fact_Suc 1);
|
|
1583 |
by (stac real_of_nat_mult 1);
|
|
1584 |
by (stac real_of_nat_mult 1);
|
|
1585 |
by (stac real_of_nat_mult 1);
|
|
1586 |
by (stac real_of_nat_mult 1);
|
12196
|
1587 |
by (simp_tac (simpset() addsimps [real_divide_def,
|
|
1588 |
real_inverse_distrib] delsimps [fact_Suc]) 1);
|
|
1589 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc RS sym]
|
|
1590 |
delsimps [fact_Suc]));
|
|
1591 |
by (multr_by_tac "real (Suc (Suc (4*m)))" 1);
|
|
1592 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc]
|
|
1593 |
delsimps [fact_Suc]));
|
|
1594 |
by (multr_by_tac "real (Suc (Suc (Suc (4*m))))" 1);
|
|
1595 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc]
|
|
1596 |
delsimps [fact_Suc]));
|
|
1597 |
by (auto_tac (claset(),simpset() addsimps [CLAIM
|
|
1598 |
"x * (x * x ^ (4*m)) = (x ^ (4*m)) * (x * (x::real))"]
|
|
1599 |
delsimps [fact_Suc]));
|
|
1600 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc,realpow_gt_zero]
|
|
1601 |
delsimps [fact_Suc]));
|
|
1602 |
by (rtac real_mult_less_mono 1);
|
|
1603 |
by (ALLGOALS(Asm_simp_tac));
|
|
1604 |
by (TRYALL(rtac real_less_trans));
|
|
1605 |
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc] delsimps [fact_Suc]));
|
|
1606 |
by (res_inst_tac [("y","0")] order_less_le_trans 1);
|
|
1607 |
by (ALLGOALS(Asm_simp_tac));
|
|
1608 |
qed "sin_gt_zero";
|
|
1609 |
|
|
1610 |
Goal "[|0 < x; x < 2 |] ==> 0 < sin x";
|
|
1611 |
by (auto_tac (claset() addIs [sin_gt_zero],simpset()));
|
|
1612 |
qed "sin_gt_zero1";
|
|
1613 |
|
|
1614 |
Goal "[| 0 < x; x < 2 |] ==> cos (2 * x) < 1";
|
|
1615 |
by (auto_tac (claset(),simpset() addsimps [cos_squared_eq,
|
|
1616 |
real_minus_add_distrib RS sym, real_minus_zero_less_iff2,sin_gt_zero1,
|
|
1617 |
real_add_order,realpow_gt_zero,cos_double] delsimps
|
|
1618 |
[realpow_Suc,real_minus_add_distrib]));
|
|
1619 |
qed "cos_double_less_one";
|
|
1620 |
|
|
1621 |
Goal "(%n. (- 1) ^ n /(real (fact (2 * n))) * x ^ (2 * n)) \
|
|
1622 |
\ sums cos x";
|
|
1623 |
by (cut_inst_tac [("x2","x")] (CLAIM "0 < (2::nat)" RS ((cos_converges
|
|
1624 |
RS sums_summable) RS sums_group)) 1);
|
|
1625 |
by (auto_tac (claset(),simpset() addsimps mult_ac@[cos_def]));
|
|
1626 |
qed "cos_paired";
|
|
1627 |
|
|
1628 |
Addsimps [realpow_gt_zero];
|
|
1629 |
|
|
1630 |
Goal "cos (2) < 0";
|
|
1631 |
by (cut_inst_tac [("x","2")] cos_paired 1);
|
|
1632 |
by (dtac sums_minus 1);
|
|
1633 |
by (rtac (CLAIM "- x < -y ==> (y::real) < x") 1);
|
|
1634 |
by (ftac sums_unique 1 THEN Auto_tac);
|
|
1635 |
by (res_inst_tac [("R2.0",
|
|
1636 |
"sumr 0 (Suc (Suc (Suc 0))) (%n. -((- 1) ^ n /(real (fact(2 * n))) \
|
|
1637 |
\ * 2 ^ (2 * n)))")] real_less_trans 1);
|
|
1638 |
by (simp_tac (simpset() addsimps [fact_num_eq_if,realpow_num_eq_if]
|
|
1639 |
delsimps [fact_Suc,realpow_Suc]) 1);
|
|
1640 |
by (simp_tac (simpset() addsimps [real_mult_assoc]
|
|
1641 |
delsimps [sumr_Suc]) 1);
|
|
1642 |
by (rtac sumr_pos_lt_pair 1);
|
|
1643 |
by (etac sums_summable 1);
|
|
1644 |
by (Step_tac 1);
|
|
1645 |
by (simp_tac (simpset() addsimps [real_divide_def,real_mult_assoc RS sym]
|
|
1646 |
delsimps [fact_Suc]) 1);
|
|
1647 |
by (rtac real_mult_inverse_cancel2 1);
|
|
1648 |
by (TRYALL(rtac (real_of_nat_fact_gt_zero)));
|
|
1649 |
by (simp_tac (simpset() addsimps [real_mult_assoc RS sym]
|
|
1650 |
delsimps [fact_Suc]) 1);
|
|
1651 |
by (rtac ((CLAIM "real(n::nat) * 4 = real(4 * n)") RS ssubst) 1);
|
12486
|
1652 |
by (stac fact_Suc 1);
|
|
1653 |
by (stac real_of_nat_mult 1);
|
|
1654 |
by (stac real_of_nat_mult 1);
|
12196
|
1655 |
by (rtac real_mult_less_mono 1);
|
|
1656 |
by (Force_tac 1);
|
|
1657 |
by (Force_tac 2);
|
|
1658 |
by (rtac real_of_nat_fact_gt_zero 2);
|
|
1659 |
by (rtac (real_of_nat_less_iff RS iffD2) 1);
|
|
1660 |
by (rtac fact_less_mono 1);
|
|
1661 |
by Auto_tac;
|
|
1662 |
qed "cos_two_less_zero";
|
|
1663 |
Addsimps [cos_two_less_zero];
|
|
1664 |
Addsimps [cos_two_less_zero RS real_not_refl2];
|
|
1665 |
Addsimps [cos_two_less_zero RS order_less_imp_le];
|
|
1666 |
|
|
1667 |
Goal "EX! x. 0 <= x & x <= 2 & cos x = 0";
|
|
1668 |
by (subgoal_tac "EX x. 0 <= x & x <= 2 & cos x = 0" 1);
|
|
1669 |
by (rtac IVT2 2);
|
|
1670 |
by (auto_tac (claset() addIs [DERIV_isCont,DERIV_cos],simpset ()));
|
|
1671 |
by (cut_inst_tac [("R1.0","xa"),("R2.0","y")] real_linear 1);
|
|
1672 |
by (rtac ccontr 1);
|
|
1673 |
by (subgoal_tac "(ALL x. cos differentiable x) & \
|
|
1674 |
\ (ALL x. isCont cos x)" 1);
|
|
1675 |
by (auto_tac (claset() addIs [DERIV_cos,DERIV_isCont],simpset()
|
|
1676 |
addsimps [differentiable_def]));
|
|
1677 |
by (dres_inst_tac [("f","cos")] Rolle 1);
|
|
1678 |
by (dres_inst_tac [("f","cos")] Rolle 5);
|
|
1679 |
by (auto_tac (claset() addSDs [DERIV_cos RS DERIV_unique],
|
|
1680 |
simpset() addsimps [differentiable_def]));
|
|
1681 |
by (dres_inst_tac [("y1","xa")] (order_le_less_trans RS sin_gt_zero) 1);
|
|
1682 |
by (assume_tac 1 THEN rtac order_less_le_trans 1);
|
|
1683 |
by (dres_inst_tac [("y1","y")] (order_le_less_trans RS sin_gt_zero) 4);
|
|
1684 |
by (assume_tac 4 THEN rtac order_less_le_trans 4);
|
|
1685 |
by (assume_tac 1 THEN assume_tac 3);
|
|
1686 |
by (ALLGOALS (Asm_full_simp_tac));
|
|
1687 |
qed "cos_is_zero";
|
|
1688 |
|
|
1689 |
Goalw [pi_def] "pi/2 = (@x. 0 <= x & x <= 2 & cos x = 0)";
|
|
1690 |
by Auto_tac;
|
|
1691 |
qed "pi_half";
|
|
1692 |
|
|
1693 |
Goal "cos (pi / 2) = 0";
|
|
1694 |
by (rtac (cos_is_zero RS ex1E) 1);
|
|
1695 |
by (auto_tac (claset() addSIs [someI2],
|
|
1696 |
simpset() addsimps [pi_half]));
|
|
1697 |
qed "cos_pi_half";
|
|
1698 |
Addsimps [cos_pi_half];
|
|
1699 |
|
|
1700 |
Goal "0 < pi / 2";
|
|
1701 |
by (rtac (cos_is_zero RS ex1E) 1);
|
|
1702 |
by (auto_tac (claset(),simpset() addsimps [pi_half]));
|
|
1703 |
by (rtac someI2 1);
|
|
1704 |
by (Blast_tac 1);
|
|
1705 |
by (Step_tac 1);
|
|
1706 |
by (dres_inst_tac [("y","xa")] real_le_imp_less_or_eq 1);
|
|
1707 |
by (Step_tac 1 THEN Asm_full_simp_tac 1);
|
|
1708 |
qed "pi_half_gt_zero";
|
|
1709 |
Addsimps [pi_half_gt_zero];
|
|
1710 |
Addsimps [(pi_half_gt_zero RS real_not_refl2) RS not_sym];
|
|
1711 |
Addsimps [pi_half_gt_zero RS order_less_imp_le];
|
|
1712 |
|
|
1713 |
Goal "pi / 2 < 2";
|
|
1714 |
by (rtac (cos_is_zero RS ex1E) 1);
|
|
1715 |
by (auto_tac (claset(),simpset() addsimps [pi_half]));
|
|
1716 |
by (rtac someI2 1);
|
|
1717 |
by (Blast_tac 1);
|
|
1718 |
by (Step_tac 1);
|
|
1719 |
by (dres_inst_tac [("x","xa")] order_le_imp_less_or_eq 1);
|
|
1720 |
by (Step_tac 1 THEN Asm_full_simp_tac 1);
|
|
1721 |
qed "pi_half_less_two";
|
|
1722 |
Addsimps [pi_half_less_two];
|
|
1723 |
Addsimps [pi_half_less_two RS real_not_refl2];
|
|
1724 |
Addsimps [pi_half_less_two RS order_less_imp_le];
|
|
1725 |
|
|
1726 |
Goal "0 < pi";
|
|
1727 |
by (multr_by_tac "inverse 2" 1);
|
|
1728 |
by Auto_tac;
|
|
1729 |
qed "pi_gt_zero";
|
|
1730 |
Addsimps [pi_gt_zero];
|
|
1731 |
Addsimps [(pi_gt_zero RS real_not_refl2) RS not_sym];
|
|
1732 |
Addsimps [pi_gt_zero RS CLAIM "(x::real) < y ==> ~ y < x"];
|
|
1733 |
|
|
1734 |
Goal "0 <= pi";
|
|
1735 |
by (auto_tac (claset() addIs [order_less_imp_le],simpset()));
|
|
1736 |
qed "pi_ge_zero";
|
|
1737 |
Addsimps [pi_ge_zero];
|
|
1738 |
|
|
1739 |
Goal "-(pi/2) < 0";
|
|
1740 |
by Auto_tac;
|
|
1741 |
qed "minus_pi_half_less_zero";
|
|
1742 |
Addsimps [minus_pi_half_less_zero];
|
|
1743 |
|
|
1744 |
Goal "sin(pi/2) = 1";
|
|
1745 |
by (cut_inst_tac [("x","pi/2")] sin_cos_squared_add2 1);
|
|
1746 |
by (cut_facts_tac [[pi_half_gt_zero,pi_half_less_two] MRS sin_gt_zero] 1);
|
|
1747 |
by (auto_tac (claset(), simpset() addsimps [numeral_2_eq_2]));
|
|
1748 |
qed "sin_pi_half";
|
|
1749 |
Addsimps [sin_pi_half];
|
|
1750 |
|
|
1751 |
Goal "cos pi = - 1";
|
|
1752 |
by (cut_inst_tac [("x","pi/2"),("y","pi/2")] cos_add 1);
|
|
1753 |
by (Asm_full_simp_tac 1);
|
|
1754 |
qed "cos_pi";
|
|
1755 |
Addsimps [cos_pi];
|
|
1756 |
|
|
1757 |
Goal "sin pi = 0";
|
|
1758 |
by (cut_inst_tac [("x","pi/2"),("y","pi/2")] sin_add 1);
|
|
1759 |
by (Asm_full_simp_tac 1);
|
|
1760 |
qed "sin_pi";
|
|
1761 |
Addsimps [sin_pi];
|
|
1762 |
|
|
1763 |
Goalw [real_diff_def] "sin x = cos (pi/2 - x)";
|
|
1764 |
by (simp_tac (simpset() addsimps [cos_add]) 1);
|
|
1765 |
qed "sin_cos_eq";
|
|
1766 |
|
|
1767 |
Goal "-sin x = cos (x + pi/2)";
|
|
1768 |
by (simp_tac (simpset() addsimps [cos_add]) 1);
|
|
1769 |
qed "minus_sin_cos_eq";
|
|
1770 |
Addsimps [minus_sin_cos_eq RS sym];
|
|
1771 |
|
|
1772 |
Goalw [real_diff_def] "cos x = sin (pi/2 - x)";
|
|
1773 |
by (simp_tac (simpset() addsimps [sin_add]) 1);
|
|
1774 |
qed "cos_sin_eq";
|
|
1775 |
Addsimps [sin_cos_eq RS sym, cos_sin_eq RS sym];
|
|
1776 |
|
|
1777 |
Goal "sin (x + pi) = - sin x";
|
|
1778 |
by (simp_tac (simpset() addsimps [sin_add]) 1);
|
|
1779 |
qed "sin_periodic_pi";
|
|
1780 |
Addsimps [sin_periodic_pi];
|
|
1781 |
|
|
1782 |
Goal "sin (pi + x) = - sin x";
|
|
1783 |
by (simp_tac (simpset() addsimps [sin_add]) 1);
|
|
1784 |
qed "sin_periodic_pi2";
|
|
1785 |
Addsimps [sin_periodic_pi2];
|
|
1786 |
|
|
1787 |
Goal "cos (x + pi) = - cos x";
|
|
1788 |
by (simp_tac (simpset() addsimps [cos_add]) 1);
|
|
1789 |
qed "cos_periodic_pi";
|
|
1790 |
Addsimps [cos_periodic_pi];
|
|
1791 |
|
|
1792 |
Goal "sin (x + 2*pi) = sin x";
|
|
1793 |
by (simp_tac (simpset() addsimps [sin_add,cos_double,numeral_2_eq_2]) 1);
|
|
1794 |
(*FIXME: just needs x^n for literals!*)
|
|
1795 |
qed "sin_periodic";
|
|
1796 |
Addsimps [sin_periodic];
|
|
1797 |
|
|
1798 |
Goal "cos (x + 2*pi) = cos x";
|
|
1799 |
by (simp_tac (simpset() addsimps [cos_add,cos_double,numeral_2_eq_2]) 1);
|
|
1800 |
(*FIXME: just needs x^n for literals!*)
|
|
1801 |
qed "cos_periodic";
|
|
1802 |
Addsimps [cos_periodic];
|
|
1803 |
|
|
1804 |
Goal "cos (real n * pi) = (-(1::real)) ^ n";
|
|
1805 |
by (induct_tac "n" 1);
|
|
1806 |
by (auto_tac (claset(),simpset() addsimps
|
|
1807 |
[real_of_nat_Suc,real_add_mult_distrib]));
|
|
1808 |
qed "cos_npi";
|
|
1809 |
Addsimps [cos_npi];
|
|
1810 |
|
|
1811 |
Goal "sin (real (n::nat) * pi) = 0";
|
|
1812 |
by (induct_tac "n" 1);
|
|
1813 |
by (auto_tac (claset(),simpset() addsimps
|
|
1814 |
[real_of_nat_Suc,real_add_mult_distrib]));
|
|
1815 |
qed "sin_npi";
|
|
1816 |
Addsimps [sin_npi];
|
|
1817 |
|
|
1818 |
Goal "sin (pi * real (n::nat)) = 0";
|
|
1819 |
by (cut_inst_tac [("n","n")] sin_npi 1);
|
|
1820 |
by (auto_tac (claset(),simpset() addsimps [real_mult_commute]
|
|
1821 |
delsimps [sin_npi]));
|
|
1822 |
qed "sin_npi2";
|
|
1823 |
Addsimps [sin_npi2];
|
|
1824 |
|
|
1825 |
Goal "cos (2 * pi) = 1";
|
|
1826 |
by (simp_tac (simpset() addsimps [cos_double,numeral_2_eq_2]) 1);
|
|
1827 |
(*FIXME: just needs x^n for literals!*)
|
|
1828 |
qed "cos_two_pi";
|
|
1829 |
Addsimps [cos_two_pi];
|
|
1830 |
|
|
1831 |
Goal "sin (2 * pi) = 0";
|
|
1832 |
by (Simp_tac 1);
|
|
1833 |
qed "sin_two_pi";
|
|
1834 |
Addsimps [sin_two_pi];
|
|
1835 |
|
|
1836 |
Goal "[| 0 < x; x < pi/2 |] ==> 0 < sin x";
|
|
1837 |
by (rtac sin_gt_zero 1);
|
|
1838 |
by (rtac real_less_trans 2 THEN assume_tac 2);
|
|
1839 |
by Auto_tac;
|
|
1840 |
qed "sin_gt_zero2";
|
|
1841 |
|
|
1842 |
Goal "[| - pi/2 < x; x < 0 |] ==> sin x < 0";
|
|
1843 |
by (rtac (CLAIM "(0::real) < - x ==> x < 0") 1);
|
|
1844 |
by (rtac (sin_minus RS subst) 1);
|
|
1845 |
by (rtac sin_gt_zero2 1);
|
|
1846 |
by (rtac (CLAIM "-y < x ==> -x < (y::real)") 2);
|
|
1847 |
by Auto_tac;
|
|
1848 |
qed "sin_less_zero";
|
|
1849 |
|
|
1850 |
Goal "[| 0 < x; x < pi/2 |] ==> 0 < cos x";
|
|
1851 |
by (cut_inst_tac [("f","cos"),("a","0"),("b","x"),("y","0")] IVT2_objl 1);
|
|
1852 |
by (Step_tac 1);
|
|
1853 |
by (cut_facts_tac [cos_is_zero] 5);
|
|
1854 |
by (Step_tac 5);
|
|
1855 |
by (dres_inst_tac [("x","xa")] spec 5);
|
|
1856 |
by (dres_inst_tac [("x","pi/2")] spec 5);
|
|
1857 |
by (auto_tac (claset() addSDs [ pi_half_less_two RS order_less_trans,
|
|
1858 |
CLAIM "~ m <= n ==> n < (m::real)"]
|
|
1859 |
addIs [DERIV_isCont,DERIV_cos],simpset()));
|
|
1860 |
qed "cos_gt_zero";
|
|
1861 |
|
|
1862 |
Goal "[| -(pi/2) < x; x < pi/2 |] ==> 0 < cos x";
|
|
1863 |
by (res_inst_tac [("R1.0","x"),("R2.0","0")] real_linear_less2 1);
|
|
1864 |
by (rtac (cos_minus RS subst) 1);
|
|
1865 |
by (rtac cos_gt_zero 1);
|
|
1866 |
by (rtac (CLAIM "-y < x ==> -x < (y::real)") 2);
|
|
1867 |
by (auto_tac (claset() addIs [cos_gt_zero],simpset()));
|
|
1868 |
qed "cos_gt_zero_pi";
|
|
1869 |
|
|
1870 |
Goal "[| -(pi/2) <= x; x <= pi/2 |] ==> 0 <= cos x";
|
|
1871 |
by (auto_tac (claset(),HOL_ss addsimps [real_le_less,
|
|
1872 |
cos_gt_zero_pi]));
|
|
1873 |
by Auto_tac;
|
|
1874 |
qed "cos_ge_zero";
|
|
1875 |
|
|
1876 |
Goal "[| 0 < x; x < pi |] ==> 0 < sin x";
|
12486
|
1877 |
by (stac sin_cos_eq 1);
|
12196
|
1878 |
by (rotate_tac 1 1);
|
|
1879 |
by (dtac (real_sum_of_halves RS ssubst) 1);
|
|
1880 |
by (auto_tac (claset() addSIs [cos_gt_zero_pi],
|
|
1881 |
simpset() delsimps [sin_cos_eq RS sym]));
|
|
1882 |
qed "sin_gt_zero_pi";
|
|
1883 |
|
|
1884 |
Goal "[| 0 <= x; x <= pi |] ==> 0 <= sin x";
|
|
1885 |
by (auto_tac (claset(),simpset() addsimps [real_le_less,
|
|
1886 |
sin_gt_zero_pi]));
|
|
1887 |
qed "sin_ge_zero";
|
|
1888 |
|
|
1889 |
Goal "[| - 1 <= y; y <= 1 |] ==> EX! x. 0 <= x & x <= pi & (cos x = y)";
|
|
1890 |
by (subgoal_tac "EX x. 0 <= x & x <= pi & cos x = y" 1);
|
|
1891 |
by (rtac IVT2 2);
|
|
1892 |
by (auto_tac (claset() addIs [order_less_imp_le,DERIV_isCont,DERIV_cos],
|
|
1893 |
simpset ()));
|
|
1894 |
by (cut_inst_tac [("R1.0","xa"),("R2.0","y")] real_linear 1);
|
|
1895 |
by (rtac ccontr 1 THEN Auto_tac);
|
|
1896 |
by (dres_inst_tac [("f","cos")] Rolle 1);
|
|
1897 |
by (dres_inst_tac [("f","cos")] Rolle 5);
|
|
1898 |
by (auto_tac (claset() addIs [order_less_imp_le,DERIV_isCont,DERIV_cos]
|
|
1899 |
addSDs [DERIV_cos RS DERIV_unique],simpset() addsimps [differentiable_def]));
|
|
1900 |
by (auto_tac (claset() addDs [[order_le_less_trans,order_less_le_trans] MRS
|
|
1901 |
sin_gt_zero_pi],simpset()));
|
|
1902 |
qed "cos_total";
|
|
1903 |
|
|
1904 |
Goal "[| - 1 <= y; y <= 1 |] ==> \
|
|
1905 |
\ EX! x. -(pi/2) <= x & x <= pi/2 & (sin x = y)";
|
|
1906 |
by (rtac ccontr 1);
|
|
1907 |
by (subgoal_tac "ALL x. (-(pi/2) <= x & x <= pi/2 & (sin x = y)) \
|
|
1908 |
\ = (0 <= (x + pi/2) & (x + pi/2) <= pi & \
|
|
1909 |
\ (cos(x + pi/2) = -y))" 1);
|
|
1910 |
by (etac swap 1);
|
|
1911 |
by (asm_full_simp_tac (simpset() delsimps [minus_sin_cos_eq RS sym]) 1);
|
|
1912 |
by (dtac (CLAIM "(x::real) <= y ==> -y <= -x") 1);
|
|
1913 |
by (dtac (CLAIM "(x::real) <= y ==> -y <= -x") 1);
|
|
1914 |
by (dtac cos_total 1);
|
|
1915 |
by (Asm_full_simp_tac 1);
|
|
1916 |
by (etac ex1E 1);
|
|
1917 |
by (res_inst_tac [("a","x - (pi/2)")] ex1I 1);
|
|
1918 |
by (simp_tac (simpset() addsimps [real_add_assoc]) 1);
|
|
1919 |
by (rotate_tac 3 1);
|
|
1920 |
by (dres_inst_tac [("x","xa + pi/2")] spec 1);
|
|
1921 |
by (Step_tac 1);
|
|
1922 |
by (TRYALL(Asm_full_simp_tac));
|
|
1923 |
by (auto_tac (claset(),simpset() addsimps [CLAIM "(-x <= y) = (-y <= (x::real))"]));
|
|
1924 |
qed "sin_total";
|
|
1925 |
|
|
1926 |
Goal "(EX n. P (n::nat)) = (EX n. P n & (ALL m. m < n --> ~ P m))";
|
|
1927 |
by (rtac iffI 1);
|
|
1928 |
by (rtac contrapos_pp 1 THEN assume_tac 1);
|
|
1929 |
by (EVERY1[Simp_tac, rtac allI, rtac nat_less_induct]);
|
|
1930 |
by (Auto_tac);
|
|
1931 |
qed "less_induct_ex_iff";
|
|
1932 |
|
|
1933 |
Goal "[| 0 < y; 0 <= x |] ==> \
|
|
1934 |
\ EX n. real n * y <= x & x < real (Suc n) * y";
|
|
1935 |
by (auto_tac (claset() addSDs [reals_Archimedean3],simpset()));
|
|
1936 |
by (dres_inst_tac [("x","x")] spec 1);
|
|
1937 |
by (dtac (less_induct_ex_iff RS iffD1) 1 THEN Step_tac 1);
|
|
1938 |
by (case_tac "n" 1);
|
|
1939 |
by (res_inst_tac [("x","nat")] exI 2);
|
|
1940 |
by Auto_tac;
|
|
1941 |
qed "reals_Archimedean4";
|
|
1942 |
|
|
1943 |
(* Pre Isabelle99-2 proof was simpler- numerals arithmetic
|
|
1944 |
now causes some unwanted re-arrangements of literals! *)
|
|
1945 |
Goal "[| 0 <= x; cos x = 0 |] ==> \
|
|
1946 |
\ EX n. ~even n & x = real n * (pi/2)";
|
|
1947 |
by (dtac (pi_gt_zero RS reals_Archimedean4) 1);
|
|
1948 |
by (Step_tac 1);
|
|
1949 |
by (subgoal_tac
|
|
1950 |
"0 <= x - real n * pi & (x - real n * pi) <= pi & \
|
|
1951 |
\ (cos(x - real n * pi) = 0)" 1);
|
|
1952 |
by (Step_tac 1);
|
|
1953 |
by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc,
|
|
1954 |
real_add_mult_distrib]) 2);
|
|
1955 |
by (asm_full_simp_tac (simpset() addsimps [cos_diff]) 1);
|
|
1956 |
by (asm_full_simp_tac (simpset() addsimps [cos_diff]) 2);
|
|
1957 |
by (subgoal_tac "EX! x. 0 <= x & x <= pi & cos x = 0" 1);
|
|
1958 |
by (rtac cos_total 2);
|
|
1959 |
by (Step_tac 1);
|
|
1960 |
by (dres_inst_tac [("x","x - real n * pi")] spec 1);
|
|
1961 |
by (dres_inst_tac [("x","pi/2")] spec 1);
|
|
1962 |
by (asm_full_simp_tac (simpset() addsimps [cos_diff]) 1);
|
|
1963 |
by (res_inst_tac [("x","Suc (2 * n)")] exI 1);
|
|
1964 |
by (asm_full_simp_tac (simpset() addsimps [real_of_nat_Suc,
|
|
1965 |
real_add_mult_distrib]) 1);
|
|
1966 |
by Auto_tac;
|
|
1967 |
qed "cos_zero_lemma";
|
|
1968 |
|
|
1969 |
Goal "[| 0 <= x; sin x = 0 |] ==> \
|
|
1970 |
\ EX n. even n & x = real n * (pi/2)";
|
|
1971 |
by (subgoal_tac
|
|
1972 |
"EX n. ~ even n & x + pi/2 = real n * (pi/2)" 1);
|
|
1973 |
by (rtac cos_zero_lemma 2);
|
|
1974 |
by (Step_tac 1);
|
|
1975 |
by (res_inst_tac [("x","n - 1")] exI 1);
|
|
1976 |
by (rtac (CLAIM "-y <= x ==> -x <= (y::real)") 2);
|
|
1977 |
by (rtac real_le_trans 2 THEN assume_tac 3);
|
|
1978 |
by (auto_tac (claset(),simpset() addsimps [odd_not_even RS sym,
|
|
1979 |
odd_Suc_mult_two_ex,real_of_nat_Suc,
|
|
1980 |
real_add_mult_distrib,real_mult_assoc RS sym]));
|
|
1981 |
qed "sin_zero_lemma";
|
|
1982 |
|
|
1983 |
(* also spoilt by numeral arithmetic *)
|
|
1984 |
Goal "(cos x = 0) = \
|
|
1985 |
\ ((EX n. ~even n & (x = real n * (pi/2))) | \
|
|
1986 |
\ (EX n. ~even n & (x = -(real n * (pi/2)))))";
|
|
1987 |
by (rtac iffI 1);
|
|
1988 |
by (cut_inst_tac [("x","x")] (CLAIM "0 <= (x::real) | x <= 0") 1);
|
|
1989 |
by (Step_tac 1);
|
|
1990 |
by (dtac cos_zero_lemma 1);
|
|
1991 |
by (dtac (CLAIM "(x::real) <= 0 ==> 0 <= -x") 3);
|
|
1992 |
by (dtac cos_zero_lemma 3);
|
|
1993 |
by (Step_tac 1);
|
|
1994 |
by (dtac (CLAIM "-x = y ==> x = -(y::real)") 2);
|
|
1995 |
by (auto_tac (claset(),HOL_ss addsimps [odd_not_even RS sym,
|
|
1996 |
odd_Suc_mult_two_ex,real_of_nat_Suc,real_add_mult_distrib]));
|
|
1997 |
by (auto_tac (claset(),simpset() addsimps [cos_add]));
|
|
1998 |
qed "cos_zero_iff";
|
|
1999 |
|
|
2000 |
(* ditto: but to a lesser extent *)
|
|
2001 |
Goal "(sin x = 0) = \
|
|
2002 |
\ ((EX n. even n & (x = real n * (pi/2))) | \
|
|
2003 |
\ (EX n. even n & (x = -(real n * (pi/2)))))";
|
|
2004 |
by (rtac iffI 1);
|
|
2005 |
by (cut_inst_tac [("x","x")] (CLAIM "0 <= (x::real) | x <= 0") 1);
|
|
2006 |
by (Step_tac 1);
|
|
2007 |
by (dtac sin_zero_lemma 1);
|
|
2008 |
by (dtac (CLAIM "(x::real) <= 0 ==> 0 <= -x") 3);
|
|
2009 |
by (dtac sin_zero_lemma 3);
|
|
2010 |
by (Step_tac 1);
|
|
2011 |
by (dtac (CLAIM "-x = y ==> x = -(y::real)") 2);
|
|
2012 |
by (auto_tac (claset(),simpset() addsimps [even_mult_two_ex]));
|
|
2013 |
qed "sin_zero_iff";
|
|
2014 |
|
|
2015 |
(* ------------------------------------------------------------------------ *)
|
|
2016 |
(* Tangent *)
|
|
2017 |
(* ------------------------------------------------------------------------ *)
|
|
2018 |
|
|
2019 |
Goalw [tan_def] "tan 0 = 0";
|
|
2020 |
by (Simp_tac 1);
|
|
2021 |
qed "tan_zero";
|
|
2022 |
Addsimps [tan_zero];
|
|
2023 |
|
|
2024 |
Goalw [tan_def] "tan pi = 0";
|
|
2025 |
by (Simp_tac 1);
|
|
2026 |
qed "tan_pi";
|
|
2027 |
Addsimps [tan_pi];
|
|
2028 |
|
|
2029 |
Goalw [tan_def] "tan (real (n::nat) * pi) = 0";
|
|
2030 |
by (Simp_tac 1);
|
|
2031 |
qed "tan_npi";
|
|
2032 |
Addsimps [tan_npi];
|
|
2033 |
|
|
2034 |
Goalw [tan_def] "tan (-x) = - tan x";
|
|
2035 |
by (simp_tac (simpset() addsimps [real_minus_mult_eq1]) 1);
|
|
2036 |
qed "tan_minus";
|
|
2037 |
Addsimps [tan_minus];
|
|
2038 |
|
|
2039 |
Goalw [tan_def] "tan (x + 2*pi) = tan x";
|
|
2040 |
by (Simp_tac 1);
|
|
2041 |
qed "tan_periodic";
|
|
2042 |
Addsimps [tan_periodic];
|
|
2043 |
|
|
2044 |
Goalw [tan_def,real_divide_def]
|
|
2045 |
"[| cos x ~= 0; cos y ~= 0 |] \
|
|
2046 |
\ ==> 1 - tan(x)*tan(y) = cos (x + y)/(cos x * cos y)";
|
|
2047 |
by (auto_tac (claset(),simpset() addsimps [real_inverse_distrib RS sym]
|
|
2048 |
@ real_mult_ac));
|
|
2049 |
by (res_inst_tac [("c1","cos x * cos y")] (real_mult_right_cancel RS subst) 1);
|
|
2050 |
by (auto_tac (claset(), simpset() addsimps [real_mult_assoc,
|
|
2051 |
real_mult_not_zero,real_diff_mult_distrib,cos_add]));
|
|
2052 |
val lemma_tan_add1 = result();
|
|
2053 |
Addsimps [lemma_tan_add1];
|
|
2054 |
|
|
2055 |
Goalw [tan_def]
|
|
2056 |
"[| cos x ~= 0; cos y ~= 0 |] \
|
|
2057 |
\ ==> tan x + tan y = sin(x + y)/(cos x * cos y)";
|
|
2058 |
by (res_inst_tac [("c1","cos x * cos y")] (real_mult_right_cancel RS subst) 1);
|
|
2059 |
by (auto_tac (claset(), simpset() addsimps [real_mult_assoc,
|
|
2060 |
real_mult_not_zero,real_add_mult_distrib]));
|
|
2061 |
by (simp_tac (simpset() addsimps [sin_add]) 1);
|
|
2062 |
qed "add_tan_eq";
|
|
2063 |
|
|
2064 |
Goal "[| cos x ~= 0; cos y ~= 0; cos (x + y) ~= 0 |] \
|
|
2065 |
\ ==> tan(x + y) = (tan(x) + tan(y))/(1 - tan(x) * tan(y))";
|
|
2066 |
by (asm_simp_tac (simpset() addsimps [add_tan_eq]) 1);
|
|
2067 |
by (simp_tac (simpset() addsimps [tan_def]) 1);
|
|
2068 |
qed "tan_add";
|
|
2069 |
|
|
2070 |
Goal "[| cos x ~= 0; cos (2 * x) ~= 0 |] \
|
|
2071 |
\ ==> tan (2 * x) = (2 * tan x)/(1 - (tan(x) ^ 2))";
|
|
2072 |
by (auto_tac (claset(),simpset() addsimps [asm_full_simplify
|
|
2073 |
(simpset() addsimps [real_mult_2 RS sym] delsimps [lemma_tan_add1])
|
|
2074 |
(read_instantiate [("x","x"),("y","x")] tan_add),numeral_2_eq_2]
|
|
2075 |
delsimps [lemma_tan_add1]));
|
|
2076 |
qed "tan_double";
|
|
2077 |
|
|
2078 |
Goalw [tan_def,real_divide_def] "[| 0 < x; x < pi/2 |] ==> 0 < tan x";
|
|
2079 |
by (auto_tac (claset() addSIs [sin_gt_zero2,cos_gt_zero_pi]
|
|
2080 |
addSIs [real_mult_order,
|
|
2081 |
real_inverse_gt_0],simpset()));
|
|
2082 |
qed "tan_gt_zero";
|
|
2083 |
|
|
2084 |
Goal "[| - pi/2 < x; x < 0 |] ==> tan x < 0";
|
|
2085 |
by (rtac (CLAIM "(0::real) < - x ==> x < 0") 1);
|
|
2086 |
by (rtac (tan_minus RS subst) 1);
|
|
2087 |
by (rtac tan_gt_zero 1);
|
|
2088 |
by (rtac (CLAIM "-x < y ==> -y < (x::real)") 2 THEN Auto_tac);
|
|
2089 |
qed "tan_less_zero";
|
|
2090 |
|
|
2091 |
Goal "cos x ~= 0 ==> DERIV (%x. sin(x)/cos(x)) x :> inverse(cos x ^ 2)";
|
|
2092 |
by (rtac lemma_DERIV_subst 1);
|
|
2093 |
by DERIV_tac;
|
|
2094 |
by (auto_tac (claset(),simpset() addsimps [real_divide_def,numeral_2_eq_2]));
|
|
2095 |
qed "lemma_DERIV_tan";
|
|
2096 |
|
|
2097 |
Goal "cos x ~= 0 ==> DERIV tan x :> inverse(cos(x) ^ 2)";
|
|
2098 |
by (auto_tac (claset() addDs [lemma_DERIV_tan],simpset()
|
|
2099 |
addsimps [(tan_def RS meta_eq_to_obj_eq) RS sym]));
|
|
2100 |
qed "DERIV_tan";
|
|
2101 |
Addsimps [DERIV_tan];
|
|
2102 |
|
|
2103 |
Goalw [real_divide_def]
|
|
2104 |
"(%x. cos(x)/sin(x)) -- pi/2 --> 0";
|
|
2105 |
by (res_inst_tac [("z1","1")] ((real_mult_0) RS subst) 1);
|
|
2106 |
by (rtac LIM_mult2 1);
|
|
2107 |
by (rtac ((real_inverse_1) RS subst) 2);
|
|
2108 |
by (rtac LIM_inverse 2);
|
|
2109 |
by (fold_tac [real_divide_def]);
|
|
2110 |
by (auto_tac (claset() addSIs [DERIV_isCont],simpset()
|
|
2111 |
addsimps [(isCont_def RS meta_eq_to_obj_eq)
|
|
2112 |
RS sym, cos_pi_half RS sym, sin_pi_half RS sym]
|
|
2113 |
delsimps [cos_pi_half,sin_pi_half]));
|
|
2114 |
by (DERIV_tac THEN Auto_tac);
|
|
2115 |
qed "LIM_cos_div_sin";
|
|
2116 |
Addsimps [LIM_cos_div_sin];
|
|
2117 |
|
|
2118 |
Goal "0 < y ==> EX x. 0 < x & x < pi/2 & y < tan x";
|
|
2119 |
by (cut_facts_tac [LIM_cos_div_sin] 1);
|
|
2120 |
by (asm_full_simp_tac (HOL_ss addsimps [LIM_def]) 1);
|
|
2121 |
by (dres_inst_tac [("x","inverse y")] spec 1);
|
|
2122 |
by (Step_tac 1);
|
|
2123 |
by (Force_tac 1);
|
|
2124 |
by (dres_inst_tac [("d1.0","s")]
|
|
2125 |
(pi_half_gt_zero RSN (2,real_lbound_gt_zero)) 1);
|
|
2126 |
by (Step_tac 1);
|
|
2127 |
by (res_inst_tac [("x","(pi/2) - e")] exI 1);
|
|
2128 |
by (Asm_simp_tac 1);
|
|
2129 |
by (dres_inst_tac [("x","(pi/2) - e")] spec 1);
|
|
2130 |
by (auto_tac (claset(),simpset() addsimps [abs_eqI2,tan_def]));
|
|
2131 |
by (rtac (real_inverse_less_iff RS iffD1) 1);
|
|
2132 |
by (auto_tac (claset(),simpset() addsimps [real_divide_def]));
|
|
2133 |
by (rtac (real_mult_order) 1);
|
|
2134 |
by (subgoal_tac "0 < sin e" 3);
|
|
2135 |
by (subgoal_tac "0 < cos e" 3);
|
|
2136 |
by (auto_tac (claset() addIs [cos_gt_zero,sin_gt_zero2],simpset()
|
|
2137 |
addsimps [real_inverse_distrib,abs_mult]));
|
|
2138 |
by (dres_inst_tac [("x","cos e")] (real_inverse_gt_0) 1);
|
|
2139 |
by (dres_inst_tac [("x","inverse (cos e)")] abs_eqI2 1);
|
|
2140 |
by (auto_tac (claset() addSDs [abs_eqI2],simpset() addsimps real_mult_ac));
|
|
2141 |
qed "lemma_tan_total";
|
|
2142 |
|
|
2143 |
|
|
2144 |
Goal "0 <= y ==> EX x. 0 <= x & x < pi/2 & tan x = y";
|
|
2145 |
by (ftac real_le_imp_less_or_eq 1);
|
|
2146 |
by (Step_tac 1 THEN Force_tac 2);
|
|
2147 |
by (dtac lemma_tan_total 1 THEN Step_tac 1);
|
|
2148 |
by (cut_inst_tac [("f","tan"),("a","0"),("b","x"),("y","y")] IVT_objl 1);
|
|
2149 |
by (auto_tac (claset() addSIs [DERIV_tan RS DERIV_isCont],simpset()));
|
|
2150 |
by (dres_inst_tac [("y","xa")] order_le_imp_less_or_eq 1);
|
|
2151 |
by (auto_tac (claset() addDs [cos_gt_zero],simpset()));
|
|
2152 |
qed "tan_total_pos";
|
|
2153 |
|
|
2154 |
Goal "EX x. -(pi/2) < x & x < (pi/2) & tan x = y";
|
|
2155 |
by (cut_inst_tac [("y","y")] (CLAIM "0 <= (y::real) | 0 <= -y") 1);
|
|
2156 |
by (Step_tac 1);
|
|
2157 |
by (dtac tan_total_pos 1);
|
|
2158 |
by (dtac tan_total_pos 2);
|
|
2159 |
by (Step_tac 1);
|
|
2160 |
by (res_inst_tac [("x","-x")] exI 2);
|
|
2161 |
by (auto_tac (claset() addSIs [exI],simpset()));
|
|
2162 |
qed "lemma_tan_total1";
|
|
2163 |
|
|
2164 |
Goal "EX! x. -(pi/2) < x & x < (pi/2) & tan x = y";
|
|
2165 |
by (cut_inst_tac [("y","y")] lemma_tan_total1 1);
|
|
2166 |
by (Auto_tac);
|
|
2167 |
by (cut_inst_tac [("R1.0","xa"),("R2.0","y")] real_linear 1);
|
|
2168 |
by (Auto_tac);
|
|
2169 |
by (subgoal_tac "EX z. xa < z & z < y & DERIV tan z :> 0" 1);
|
|
2170 |
by (subgoal_tac "EX z. y < z & z < xa & DERIV tan z :> 0" 3);
|
|
2171 |
by (rtac Rolle 2);
|
|
2172 |
by (rtac Rolle 7);
|
|
2173 |
by (auto_tac (claset() addSIs [DERIV_tan,DERIV_isCont,exI],simpset()
|
|
2174 |
addsimps [differentiable_def]));
|
|
2175 |
by (TRYALL(rtac DERIV_tan));
|
|
2176 |
by (TRYALL(dtac (DERIV_tan RSN (2,DERIV_unique))));
|
|
2177 |
by (TRYALL(rtac (real_not_refl2 RS not_sym)));
|
|
2178 |
by (auto_tac (claset() addSIs [cos_gt_zero_pi],simpset()));
|
|
2179 |
by (ALLGOALS(subgoal_tac "0 < cos z"));
|
|
2180 |
by (Force_tac 1 THEN Force_tac 2);
|
|
2181 |
by (ALLGOALS(thin_tac "cos z = 0"));
|
|
2182 |
by (auto_tac (claset() addSIs [cos_gt_zero_pi],simpset()));
|
|
2183 |
qed "tan_total";
|
|
2184 |
|
|
2185 |
Goal "[| - 1 <= y; y <= 1 |] \
|
|
2186 |
\ ==> -(pi/2) <= arcsin y & arcsin y <= pi & sin(arcsin y) = y";
|
|
2187 |
by (dtac sin_total 1);
|
|
2188 |
by (etac ex1E 2);
|
|
2189 |
by (rewtac arcsin_def);
|
|
2190 |
by (rtac someI2 2);
|
|
2191 |
by (EVERY1[assume_tac, Blast_tac, Step_tac]);
|
|
2192 |
by (rtac real_le_trans 1 THEN assume_tac 1);
|
|
2193 |
by (Force_tac 1);
|
|
2194 |
qed "arcsin_pi";
|
|
2195 |
|
|
2196 |
Goal "[| - 1 <= y; y <= 1 |] \
|
|
2197 |
\ ==> -(pi/2) <= arcsin y & \
|
|
2198 |
\ arcsin y <= pi/2 & sin(arcsin y) = y";
|
|
2199 |
by (dtac sin_total 1 THEN assume_tac 1);
|
|
2200 |
by (etac ex1E 1);
|
|
2201 |
by (rewtac arcsin_def);
|
|
2202 |
by (rtac someI2 1);
|
|
2203 |
by (ALLGOALS(Blast_tac));
|
|
2204 |
qed "arcsin";
|
|
2205 |
|
|
2206 |
Goal "[| - 1 <= y; y <= 1 |] ==> sin(arcsin y) = y";
|
|
2207 |
by (blast_tac (claset() addDs [arcsin]) 1);
|
|
2208 |
qed "sin_arcsin";
|
|
2209 |
Addsimps [sin_arcsin];
|
|
2210 |
|
|
2211 |
Goal "[| -1 <= y; y <= 1 |] ==> sin(arcsin y) = y";
|
|
2212 |
by (auto_tac (claset() addIs [sin_arcsin],simpset()));
|
|
2213 |
qed "sin_arcsin2";
|
|
2214 |
Addsimps [sin_arcsin2];
|
|
2215 |
|
|
2216 |
Goal "[| - 1 <= y; y <= 1 |] \
|
|
2217 |
\ ==> -(pi/2) <= arcsin y & arcsin y <= pi/2";
|
|
2218 |
by (blast_tac (claset() addDs [arcsin]) 1);
|
|
2219 |
qed "arcsin_bounded";
|
|
2220 |
|
|
2221 |
Goal "[| - 1 <= y; y <= 1 |] ==> -(pi/2) <= arcsin y";
|
|
2222 |
by (blast_tac (claset() addDs [arcsin]) 1);
|
|
2223 |
qed "arcsin_lbound";
|
|
2224 |
|
|
2225 |
Goal "[| - 1 <= y; y <= 1 |] ==> arcsin y <= pi/2";
|
|
2226 |
by (blast_tac (claset() addDs [arcsin]) 1);
|
|
2227 |
qed "arcsin_ubound";
|
|
2228 |
|
|
2229 |
Goal "[| - 1 < y; y < 1 |] \
|
|
2230 |
\ ==> -(pi/2) < arcsin y & arcsin y < pi/2";
|
|
2231 |
by (ftac order_less_imp_le 1);
|
|
2232 |
by (forw_inst_tac [("y","y")] order_less_imp_le 1);
|
|
2233 |
by (ftac arcsin_bounded 1);
|
|
2234 |
by (Step_tac 1 THEN Asm_full_simp_tac 1);
|
|
2235 |
by (dres_inst_tac [("y","arcsin y")] order_le_imp_less_or_eq 1);
|
|
2236 |
by (dres_inst_tac [("y","pi/2")] order_le_imp_less_or_eq 2);
|
|
2237 |
by (Step_tac 1);
|
|
2238 |
by (ALLGOALS(dres_inst_tac [("f","sin")] arg_cong));
|
|
2239 |
by (Auto_tac);
|
|
2240 |
qed "arcsin_lt_bounded";
|
|
2241 |
|
|
2242 |
Goalw [arcsin_def]
|
|
2243 |
"[|-(pi/2) <= x; x <= pi/2 |] ==> arcsin(sin x) = x";
|
|
2244 |
by (rtac some1_equality 1);
|
|
2245 |
by (rtac sin_total 1);
|
|
2246 |
by Auto_tac;
|
|
2247 |
qed "arcsin_sin";
|
|
2248 |
|
|
2249 |
Goal "[| - 1 <= y; y <= 1 |] \
|
|
2250 |
\ ==> 0 <= arcos y & arcos y <= pi & cos(arcos y) = y";
|
|
2251 |
by (dtac cos_total 1 THEN assume_tac 1);
|
|
2252 |
by (etac ex1E 1);
|
|
2253 |
by (rewtac arcos_def);
|
|
2254 |
by (rtac someI2 1);
|
|
2255 |
by (ALLGOALS(Blast_tac));
|
|
2256 |
qed "arcos";
|
|
2257 |
|
|
2258 |
Goal "[| - 1 <= y; y <= 1 |] ==> cos(arcos y) = y";
|
|
2259 |
by (blast_tac (claset() addDs [arcos]) 1);
|
|
2260 |
qed "cos_arcos";
|
|
2261 |
Addsimps [cos_arcos];
|
|
2262 |
|
|
2263 |
Goal "[| -1 <= y; y <= 1 |] ==> cos(arcos y) = y";
|
|
2264 |
by (auto_tac (claset() addIs [cos_arcos],simpset()));
|
|
2265 |
qed "cos_arcos2";
|
|
2266 |
Addsimps [cos_arcos2];
|
|
2267 |
|
|
2268 |
Goal "[| - 1 <= y; y <= 1 |] ==> 0 <= arcos y & arcos y <= pi";
|
|
2269 |
by (blast_tac (claset() addDs [arcos]) 1);
|
|
2270 |
qed "arcos_bounded";
|
|
2271 |
|
|
2272 |
Goal "[| - 1 <= y; y <= 1 |] ==> 0 <= arcos y";
|
|
2273 |
by (blast_tac (claset() addDs [arcos]) 1);
|
|
2274 |
qed "arcos_lbound";
|
|
2275 |
|
|
2276 |
Goal "[| - 1 <= y; y <= 1 |] ==> arcos y <= pi";
|
|
2277 |
by (blast_tac (claset() addDs [arcos]) 1);
|
|
2278 |
qed "arcos_ubound";
|
|
2279 |
|
|
2280 |
Goal "[| - 1 < y; y < 1 |] \
|
|
2281 |
\ ==> 0 < arcos y & arcos y < pi";
|
|
2282 |
by (ftac order_less_imp_le 1);
|
|
2283 |
by (forw_inst_tac [("y","y")] order_less_imp_le 1);
|
|
2284 |
by (ftac arcos_bounded 1);
|
|
2285 |
by (Auto_tac);
|
|
2286 |
by (dres_inst_tac [("y","arcos y")] order_le_imp_less_or_eq 1);
|
|
2287 |
by (dres_inst_tac [("y","pi")] order_le_imp_less_or_eq 2);
|
|
2288 |
by (Auto_tac);
|
|
2289 |
by (ALLGOALS(dres_inst_tac [("f","cos")] arg_cong));
|
|
2290 |
by (Auto_tac);
|
|
2291 |
qed "arcos_lt_bounded";
|
|
2292 |
|
|
2293 |
Goalw [arcos_def] "[|0 <= x; x <= pi |] ==> arcos(cos x) = x";
|
|
2294 |
by (auto_tac (claset() addSIs [some1_equality,cos_total],simpset()));
|
|
2295 |
qed "arcos_cos";
|
|
2296 |
|
|
2297 |
Goalw [arcos_def] "[|x <= 0; -pi <= x |] ==> arcos(cos x) = -x";
|
|
2298 |
by (auto_tac (claset() addSIs [some1_equality,cos_total],simpset()));
|
|
2299 |
qed "arcos_cos2";
|
|
2300 |
|
|
2301 |
Goal "- (pi/2) < arctan y & \
|
|
2302 |
\ arctan y < pi/2 & tan (arctan y) = y";
|
|
2303 |
by (cut_inst_tac [("y","y")] tan_total 1);
|
|
2304 |
by (etac ex1E 1);
|
|
2305 |
by (rewtac arctan_def);
|
|
2306 |
by (rtac someI2 1);
|
|
2307 |
by (ALLGOALS(Blast_tac));
|
|
2308 |
qed "arctan";
|
|
2309 |
Addsimps [arctan];
|
|
2310 |
|
|
2311 |
Goal "tan(arctan y) = y";
|
|
2312 |
by (Auto_tac);
|
|
2313 |
qed "tan_arctan";
|
|
2314 |
|
|
2315 |
Goal "- (pi/2) < arctan y & arctan y < pi/2";
|
|
2316 |
by (Auto_tac);
|
|
2317 |
qed "arctan_bounded";
|
|
2318 |
|
|
2319 |
Goal "- (pi/2) < arctan y";
|
|
2320 |
by (Auto_tac);
|
|
2321 |
qed "arctan_lbound";
|
|
2322 |
|
|
2323 |
Goal "arctan y < pi/2";
|
|
2324 |
by (Auto_tac);
|
|
2325 |
qed "arctan_ubound";
|
|
2326 |
|
|
2327 |
Goalw [arctan_def]
|
|
2328 |
"[|-(pi/2) < x; x < pi/2 |] ==> arctan(tan x) = x";
|
|
2329 |
by (rtac some1_equality 1);
|
|
2330 |
by (rtac tan_total 1);
|
|
2331 |
by Auto_tac;
|
|
2332 |
qed "arctan_tan";
|
|
2333 |
|
|
2334 |
Goal "arctan 0 = 0";
|
|
2335 |
by (rtac (asm_full_simplify (simpset())
|
|
2336 |
(read_instantiate [("x","0")] arctan_tan)) 1);
|
|
2337 |
qed "arctan_zero_zero";
|
|
2338 |
Addsimps [arctan_zero_zero];
|
|
2339 |
|
|
2340 |
(* ------------------------------------------------------------------------- *)
|
|
2341 |
(* Differentiation of arctan. *)
|
|
2342 |
(* ------------------------------------------------------------------------- *)
|
|
2343 |
|
|
2344 |
Goal "cos(arctan x) ~= 0";
|
|
2345 |
by (auto_tac (claset(),simpset() addsimps [cos_zero_iff]));
|
|
2346 |
by (case_tac "n" 1);
|
|
2347 |
by (case_tac "n" 3);
|
|
2348 |
by (cut_inst_tac [("y","x")] arctan_ubound 2);
|
|
2349 |
by (cut_inst_tac [("y","x")] arctan_lbound 4);
|
|
2350 |
by (auto_tac (claset(),simpset() addsimps [real_of_nat_Suc,
|
|
2351 |
real_add_mult_distrib,real_le_def,
|
|
2352 |
real_mult_less_0_iff] delsimps [arctan]));
|
|
2353 |
qed "cos_arctan_not_zero";
|
|
2354 |
Addsimps [cos_arctan_not_zero];
|
|
2355 |
|
|
2356 |
Goal "cos x ~= 0 ==> 1 + tan(x) ^ 2 = inverse(cos x) ^ 2";
|
|
2357 |
by (rtac (realpow_inverse RS subst) 1);
|
|
2358 |
by (res_inst_tac [("c1","cos(x) ^ 2")] (real_mult_right_cancel RS iffD1) 1);
|
|
2359 |
by (auto_tac (claset() addDs [realpow_not_zero], simpset() addsimps
|
|
2360 |
[realpow_mult,real_add_mult_distrib,realpow_divide,
|
|
2361 |
tan_def,real_mult_assoc,realpow_inverse RS sym]
|
|
2362 |
delsimps [realpow_Suc]));
|
|
2363 |
qed "tan_sec";
|
|
2364 |
|
|
2365 |
|
|
2366 |
(*--------------------------------------------------------------------------*)
|
|
2367 |
(* Some more theorems- developed while at ICASE (07/2001) *)
|
|
2368 |
(*--------------------------------------------------------------------------*)
|
|
2369 |
|
|
2370 |
Goal "sin (xa + 1 / 2 * real (Suc m) * pi) = \
|
|
2371 |
\ cos (xa + 1 / 2 * real (m) * pi)";
|
|
2372 |
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,
|
|
2373 |
real_of_nat_Suc,real_add_mult_distrib,real_add_mult_distrib2]) 1);
|
|
2374 |
by Auto_tac;
|
|
2375 |
qed "lemma_sin_cos_eq";
|
|
2376 |
Addsimps [lemma_sin_cos_eq];
|
|
2377 |
|
|
2378 |
Goal "sin (xa + real (Suc m) * pi / 2) = \
|
|
2379 |
\ cos (xa + real (m) * pi / 2)";
|
|
2380 |
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,real_divide_def,
|
|
2381 |
real_of_nat_Suc,real_add_mult_distrib,real_add_mult_distrib2]) 1);
|
|
2382 |
by Auto_tac;
|
|
2383 |
qed "lemma_sin_cos_eq2";
|
|
2384 |
Addsimps [lemma_sin_cos_eq2];
|
|
2385 |
|
|
2386 |
Goal "DERIV (%x. sin (x + k)) xa :> cos (xa + k)";
|
|
2387 |
by (rtac lemma_DERIV_subst 1);
|
|
2388 |
by (res_inst_tac [("f","sin"),("g","%x. x + k")] DERIV_chain2 1);
|
|
2389 |
by DERIV_tac;
|
|
2390 |
by (Simp_tac 1);
|
|
2391 |
qed "DERIV_sin_add";
|
|
2392 |
Addsimps [DERIV_sin_add];
|
|
2393 |
|
|
2394 |
(* which further simplifies to (- 1 ^ m) !! *)
|
|
2395 |
Goal "sin ((real m + 1/2) * pi) = cos (real m * pi)";
|
|
2396 |
by (auto_tac (claset(),simpset() addsimps [real_add_mult_distrib2,
|
|
2397 |
sin_add,real_add_mult_distrib] @ real_mult_ac));
|
|
2398 |
qed "sin_cos_npi";
|
|
2399 |
Addsimps [sin_cos_npi];
|
|
2400 |
|
|
2401 |
Goal "sin (real (Suc (2 * n)) * pi / 2) = (- 1) ^ n";
|
|
2402 |
by (cut_inst_tac [("m","n")] sin_cos_npi 1);
|
|
2403 |
by (auto_tac (claset(),HOL_ss addsimps [real_of_nat_Suc,
|
|
2404 |
real_add_mult_distrib,real_divide_def]));
|
|
2405 |
by Auto_tac;
|
|
2406 |
qed "sin_cos_npi2";
|
|
2407 |
Addsimps [ sin_cos_npi2];
|
|
2408 |
|
|
2409 |
Goal "cos (2 * real (n::nat) * pi) = 1";
|
|
2410 |
by (auto_tac (claset(),simpset() addsimps [cos_double,
|
|
2411 |
real_mult_assoc,realpow_add RS sym,numeral_2_eq_2]));
|
|
2412 |
(*FIXME: just needs x^n for literals!*)
|
|
2413 |
qed "cos_2npi";
|
|
2414 |
Addsimps [cos_2npi];
|
|
2415 |
|
|
2416 |
Goal "cos (3 / 2 * pi) = 0";
|
|
2417 |
by (rtac (CLAIM "(1::real) + 1/2 = 3/2" RS subst) 1);
|
12486
|
2418 |
by (stac real_add_mult_distrib 1);
|
12196
|
2419 |
by (auto_tac (claset(),simpset() addsimps [cos_add] @ real_mult_ac));
|
|
2420 |
qed "cos_3over2_pi";
|
|
2421 |
Addsimps [cos_3over2_pi];
|
|
2422 |
|
|
2423 |
Goal "sin (2 * real (n::nat) * pi) = 0";
|
|
2424 |
by (auto_tac (claset(),simpset() addsimps [real_mult_assoc]));
|
|
2425 |
qed "sin_2npi";
|
|
2426 |
Addsimps [sin_2npi];
|
|
2427 |
|
|
2428 |
Goal "sin (3 / 2 * pi) = - 1";
|
|
2429 |
by (rtac (CLAIM "(1::real) + 1/2 = 3/2" RS subst) 1);
|
12486
|
2430 |
by (stac real_add_mult_distrib 1);
|
12196
|
2431 |
by (auto_tac (claset(),simpset() addsimps [sin_add] @real_mult_ac));
|
|
2432 |
qed "sin_3over2_pi";
|
|
2433 |
Addsimps [sin_3over2_pi];
|
|
2434 |
|
|
2435 |
Goal "cos(xa + 1 / 2 * real (Suc m) * pi) = \
|
|
2436 |
\ -sin (xa + 1 / 2 * real (m) * pi)";
|
|
2437 |
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,
|
|
2438 |
real_of_nat_Suc,real_add_mult_distrib2,real_add_mult_distrib,
|
|
2439 |
real_minus_mult_eq2]) 1);
|
|
2440 |
by Auto_tac;
|
|
2441 |
qed "lemma_cos_sin_eq";
|
|
2442 |
Addsimps [lemma_cos_sin_eq];
|
|
2443 |
|
|
2444 |
Goal "cos (xa + real (Suc m) * pi / 2) = \
|
|
2445 |
\ -sin (xa + real (m) * pi / 2)";
|
|
2446 |
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,real_divide_def,
|
|
2447 |
real_of_nat_Suc,real_add_mult_distrib,real_add_mult_distrib2]) 1);
|
|
2448 |
by Auto_tac;
|
|
2449 |
qed "lemma_cos_sin_eq2";
|
|
2450 |
Addsimps [lemma_cos_sin_eq2];
|
|
2451 |
|
|
2452 |
Goal "cos (pi * real (Suc (2 * m)) / 2) = 0";
|
|
2453 |
by (simp_tac (HOL_ss addsimps [cos_add,sin_add,real_divide_def,
|
|
2454 |
real_of_nat_Suc,real_add_mult_distrib,real_add_mult_distrib2]) 1);
|
|
2455 |
by Auto_tac;
|
|
2456 |
qed "cos_pi_eq_zero";
|
|
2457 |
Addsimps [cos_pi_eq_zero];
|
|
2458 |
|
|
2459 |
Goal "DERIV (%x. cos (x + k)) xa :> - sin (xa + k)";
|
|
2460 |
by (rtac lemma_DERIV_subst 1);
|
|
2461 |
by (res_inst_tac [("f","cos"),("g","%x. x + k")] DERIV_chain2 1);
|
|
2462 |
by DERIV_tac;
|
|
2463 |
by (Simp_tac 1);
|
|
2464 |
qed "DERIV_cos_add";
|
|
2465 |
Addsimps [DERIV_cos_add];
|
|
2466 |
|
|
2467 |
Goal "isCont cos x";
|
|
2468 |
by (rtac (DERIV_cos RS DERIV_isCont) 1);
|
|
2469 |
qed "isCont_cos";
|
|
2470 |
|
|
2471 |
Goal "isCont sin x";
|
|
2472 |
by (rtac (DERIV_sin RS DERIV_isCont) 1);
|
|
2473 |
qed "isCont_sin";
|
|
2474 |
|
|
2475 |
Goal "isCont exp x";
|
|
2476 |
by (rtac (DERIV_exp RS DERIV_isCont) 1);
|
|
2477 |
qed "isCont_exp";
|
|
2478 |
|
|
2479 |
val isCont_simp = [isCont_exp,isCont_sin,isCont_cos];
|
|
2480 |
Addsimps isCont_simp;
|
|
2481 |
|
|
2482 |
(** more theorems: e.g. used in complex geometry **)
|
|
2483 |
|
|
2484 |
Goal "sin x = 0 ==> abs(cos x) = 1";
|
|
2485 |
by (auto_tac (claset(),simpset() addsimps [sin_zero_iff,even_mult_two_ex]));
|
|
2486 |
qed "sin_zero_abs_cos_one";
|
|
2487 |
|
|
2488 |
Goal "(exp x = 1) = (x = 0)";
|
|
2489 |
by Auto_tac;
|
|
2490 |
by (dres_inst_tac [("f","ln")] arg_cong 1);
|
|
2491 |
by Auto_tac;
|
|
2492 |
qed "exp_eq_one_iff";
|
|
2493 |
Addsimps [exp_eq_one_iff];
|
|
2494 |
|
|
2495 |
Goal "cos x = 1 ==> sin x = 0";
|
|
2496 |
by (cut_inst_tac [("x","x")] sin_cos_squared_add3 1);
|
|
2497 |
by Auto_tac;
|
|
2498 |
qed "cos_one_sin_zero";
|
|
2499 |
|
|
2500 |
(*-------------------------------------------------------------------------------*)
|
|
2501 |
(* Add continuity checker in backup of theory? *)
|
|
2502 |
(*-------------------------------------------------------------------------------*)
|
|
2503 |
|