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