| author | blanchet | 
| Thu, 29 Apr 2010 19:07:36 +0200 | |
| changeset 36567 | f1cb249f6384 | 
| parent 36310 | e3d3b14b13cd | 
| child 37765 | 26bdfb7b680b | 
| permissions | -rw-r--r-- | 
| 27468 | 1 | (* Title : Deriv.thy | 
| 2 | Author : Jacques D. Fleuriot | |
| 3 | Copyright : 1998 University of Cambridge | |
| 4 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | |
| 5 | *) | |
| 6 | ||
| 7 | header{* Differentiation (Nonstandard) *}
 | |
| 8 | ||
| 9 | theory HDeriv | |
| 10 | imports Deriv HLim | |
| 11 | begin | |
| 12 | ||
| 13 | text{*Nonstandard Definitions*}
 | |
| 14 | ||
| 15 | definition | |
| 16 | nsderiv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool" | |
| 17 |           ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60) where
 | |
| 18 |   "NSDERIV f x :> D = (\<forall>h \<in> Infinitesimal - {0}.
 | |
| 19 | (( *f* f)(star_of x + h) | |
| 20 | - star_of (f x))/h @= star_of D)" | |
| 21 | ||
| 22 | definition | |
| 23 | NSdifferentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool" | |
| 24 | (infixl "NSdifferentiable" 60) where | |
| 25 | "f NSdifferentiable x = (\<exists>D. NSDERIV f x :> D)" | |
| 26 | ||
| 27 | definition | |
| 28 | increment :: "[real=>real,real,hypreal] => hypreal" where | |
| 28562 | 29 | [code del]: "increment f x h = (@inc. f NSdifferentiable x & | 
| 27468 | 30 | inc = ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x))" | 
| 31 | ||
| 32 | ||
| 33 | subsection {* Derivatives *}
 | |
| 34 | ||
| 35 | lemma DERIV_NS_iff: | |
| 36 | "(DERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NS> D)" | |
| 37 | by (simp add: deriv_def LIM_NSLIM_iff) | |
| 38 | ||
| 39 | lemma NS_DERIV_D: "DERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --NS> D" | |
| 40 | by (simp add: deriv_def LIM_NSLIM_iff) | |
| 41 | ||
| 42 | lemma hnorm_of_hypreal: | |
| 43 | "\<And>r. hnorm (( *f* of_real) r::'a::real_normed_div_algebra star) = \<bar>r\<bar>" | |
| 44 | by transfer (rule norm_of_real) | |
| 45 | ||
| 46 | lemma Infinitesimal_of_hypreal: | |
| 47 | "x \<in> Infinitesimal \<Longrightarrow> | |
| 48 | (( *f* of_real) x::'a::real_normed_div_algebra star) \<in> Infinitesimal" | |
| 49 | apply (rule InfinitesimalI2) | |
| 50 | apply (drule (1) InfinitesimalD2) | |
| 51 | apply (simp add: hnorm_of_hypreal) | |
| 52 | done | |
| 53 | ||
| 54 | lemma of_hypreal_eq_0_iff: | |
| 55 | "\<And>x. (( *f* of_real) x = (0::'a::real_algebra_1 star)) = (x = 0)" | |
| 56 | by transfer (rule of_real_eq_0_iff) | |
| 57 | ||
| 58 | lemma NSDeriv_unique: | |
| 59 | "[| NSDERIV f x :> D; NSDERIV f x :> E |] ==> D = E" | |
| 60 | apply (subgoal_tac "( *f* of_real) epsilon \<in> Infinitesimal - {0::'a star}")
 | |
| 61 | apply (simp only: nsderiv_def) | |
| 62 | apply (drule (1) bspec)+ | |
| 63 | apply (drule (1) approx_trans3) | |
| 64 | apply simp | |
| 65 | apply (simp add: Infinitesimal_of_hypreal Infinitesimal_epsilon) | |
| 66 | apply (simp add: of_hypreal_eq_0_iff hypreal_epsilon_not_zero) | |
| 67 | done | |
| 68 | ||
| 69 | text {*First NSDERIV in terms of NSLIM*}
 | |
| 70 | ||
| 71 | text{*first equivalence *}
 | |
| 72 | lemma NSDERIV_NSLIM_iff: | |
| 73 | "(NSDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NS> D)" | |
| 74 | apply (simp add: nsderiv_def NSLIM_def, auto) | |
| 75 | apply (drule_tac x = xa in bspec) | |
| 76 | apply (rule_tac [3] ccontr) | |
| 77 | apply (drule_tac [3] x = h in spec) | |
| 78 | apply (auto simp add: mem_infmal_iff starfun_lambda_cancel) | |
| 79 | done | |
| 80 | ||
| 81 | text{*second equivalence *}
 | |
| 82 | lemma NSDERIV_NSLIM_iff2: | |
| 83 | "(NSDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z-x)) -- x --NS> D)" | |
| 84 | by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff diff_minus [symmetric] | |
| 85 | LIM_NSLIM_iff [symmetric]) | |
| 86 | ||
| 87 | (* while we're at it! *) | |
| 88 | ||
| 89 | lemma NSDERIV_iff2: | |
| 90 | "(NSDERIV f x :> D) = | |
| 91 | (\<forall>w. | |
| 92 | w \<noteq> star_of x & w \<approx> star_of x --> | |
| 93 | ( *f* (%z. (f z - f x) / (z-x))) w \<approx> star_of D)" | |
| 94 | by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def) | |
| 95 | ||
| 96 | (*FIXME DELETE*) | |
| 97 | lemma hypreal_not_eq_minus_iff: | |
| 98 | "(x \<noteq> a) = (x - a \<noteq> (0::'a::ab_group_add))" | |
| 99 | by auto | |
| 100 | ||
| 101 | lemma NSDERIVD5: | |
| 102 | "(NSDERIV f x :> D) ==> | |
| 103 | (\<forall>u. u \<approx> hypreal_of_real x --> | |
| 104 | ( *f* (%z. f z - f x)) u \<approx> hypreal_of_real D * (u - hypreal_of_real x))" | |
| 105 | apply (auto simp add: NSDERIV_iff2) | |
| 106 | apply (case_tac "u = hypreal_of_real x", auto) | |
| 107 | apply (drule_tac x = u in spec, auto) | |
| 108 | apply (drule_tac c = "u - hypreal_of_real x" and b = "hypreal_of_real D" in approx_mult1) | |
| 109 | apply (drule_tac [!] hypreal_not_eq_minus_iff [THEN iffD1]) | |
| 110 | apply (subgoal_tac [2] "( *f* (%z. z-x)) u \<noteq> (0::hypreal) ") | |
| 111 | apply (auto simp add: | |
| 112 | approx_minus_iff [THEN iffD1, THEN mem_infmal_iff [THEN iffD2]] | |
| 113 | Infinitesimal_subset_HFinite [THEN subsetD]) | |
| 114 | done | |
| 115 | ||
| 116 | lemma NSDERIVD4: | |
| 117 | "(NSDERIV f x :> D) ==> | |
| 118 | (\<forall>h \<in> Infinitesimal. | |
| 119 | (( *f* f)(hypreal_of_real x + h) - | |
| 120 | hypreal_of_real (f x))\<approx> (hypreal_of_real D) * h)" | |
| 121 | apply (auto simp add: nsderiv_def) | |
| 122 | apply (case_tac "h = (0::hypreal) ") | |
| 123 | apply (auto simp add: diff_minus) | |
| 124 | apply (drule_tac x = h in bspec) | |
| 125 | apply (drule_tac [2] c = h in approx_mult1) | |
| 126 | apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] | |
| 127 | simp add: diff_minus) | |
| 128 | done | |
| 129 | ||
| 130 | lemma NSDERIVD3: | |
| 131 | "(NSDERIV f x :> D) ==> | |
| 132 |       (\<forall>h \<in> Infinitesimal - {0}.
 | |
| 133 | (( *f* f)(hypreal_of_real x + h) - | |
| 134 | hypreal_of_real (f x))\<approx> (hypreal_of_real D) * h)" | |
| 135 | apply (auto simp add: nsderiv_def) | |
| 136 | apply (rule ccontr, drule_tac x = h in bspec) | |
| 137 | apply (drule_tac [2] c = h in approx_mult1) | |
| 138 | apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] | |
| 139 | simp add: mult_assoc diff_minus) | |
| 140 | done | |
| 141 | ||
| 142 | text{*Differentiability implies continuity
 | |
| 143 | nice and simple "algebraic" proof*} | |
| 144 | lemma NSDERIV_isNSCont: "NSDERIV f x :> D ==> isNSCont f x" | |
| 145 | apply (auto simp add: nsderiv_def isNSCont_NSLIM_iff NSLIM_def) | |
| 146 | apply (drule approx_minus_iff [THEN iffD1]) | |
| 147 | apply (drule hypreal_not_eq_minus_iff [THEN iffD1]) | |
| 148 | apply (drule_tac x = "xa - star_of x" in bspec) | |
| 149 | prefer 2 apply (simp add: add_assoc [symmetric]) | |
| 150 | apply (auto simp add: mem_infmal_iff [symmetric] add_commute) | |
| 151 | apply (drule_tac c = "xa - star_of x" in approx_mult1) | |
| 152 | apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] | |
| 153 | simp add: mult_assoc nonzero_mult_divide_cancel_right) | |
| 154 | apply (drule_tac x3=D in | |
| 155 | HFinite_star_of [THEN [2] Infinitesimal_HFinite_mult, | |
| 156 | THEN mem_infmal_iff [THEN iffD1]]) | |
| 157 | apply (auto simp add: mult_commute | |
| 158 | intro: approx_trans approx_minus_iff [THEN iffD2]) | |
| 159 | done | |
| 160 | ||
| 161 | text{*Differentiation rules for combinations of functions
 | |
| 162 | follow from clear, straightforard, algebraic | |
| 163 | manipulations*} | |
| 164 | text{*Constant function*}
 | |
| 165 | ||
| 166 | (* use simple constant nslimit theorem *) | |
| 167 | lemma NSDERIV_const [simp]: "(NSDERIV (%x. k) x :> 0)" | |
| 168 | by (simp add: NSDERIV_NSLIM_iff) | |
| 169 | ||
| 170 | text{*Sum of functions- proved easily*}
 | |
| 171 | ||
| 172 | lemma NSDERIV_add: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] | |
| 173 | ==> NSDERIV (%x. f x + g x) x :> Da + Db" | |
| 174 | apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) | |
| 175 | apply (auto simp add: add_divide_distrib diff_divide_distrib dest!: spec) | |
| 176 | apply (drule_tac b = "star_of Da" and d = "star_of Db" in approx_add) | |
| 177 | apply (auto simp add: diff_def add_ac) | |
| 178 | done | |
| 179 | ||
| 180 | text{*Product of functions - Proof is trivial but tedious
 | |
| 181 | and long due to rearrangement of terms*} | |
| 182 | ||
| 183 | lemma lemma_nsderiv1: | |
| 184 | fixes a b c d :: "'a::comm_ring star" | |
| 185 | shows "(a*b) - (c*d) = (b*(a - c)) + (c*(b - d))" | |
| 186 | by (simp add: right_diff_distrib mult_ac) | |
| 187 | ||
| 188 | lemma lemma_nsderiv2: | |
| 189 | fixes x y z :: "'a::real_normed_field star" | |
| 190 | shows "[| (x - y) / z = star_of D + yb; z \<noteq> 0; | |
| 191 | z \<in> Infinitesimal; yb \<in> Infinitesimal |] | |
| 192 | ==> x - y \<approx> 0" | |
| 193 | apply (simp add: nonzero_divide_eq_eq) | |
| 194 | apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add | |
| 195 | simp add: mult_assoc mem_infmal_iff [symmetric]) | |
| 196 | apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) | |
| 197 | done | |
| 198 | ||
| 199 | lemma NSDERIV_mult: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] | |
| 200 | ==> NSDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" | |
| 201 | apply (auto simp add: NSDERIV_NSLIM_iff NSLIM_def) | |
| 202 | apply (auto dest!: spec | |
| 203 | simp add: starfun_lambda_cancel lemma_nsderiv1) | |
| 204 | apply (simp (no_asm) add: add_divide_distrib diff_divide_distrib) | |
| 205 | apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+ | |
| 206 | apply (auto simp add: times_divide_eq_right [symmetric] | |
| 36310 | 207 | simp del: times_divide_eq_right times_divide_eq_left) | 
| 27468 | 208 | apply (drule_tac D = Db in lemma_nsderiv2, assumption+) | 
| 209 | apply (drule_tac | |
| 210 | approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]]) | |
| 211 | apply (auto intro!: approx_add_mono1 | |
| 212 | simp add: left_distrib right_distrib mult_commute add_assoc) | |
| 213 | apply (rule_tac b1 = "star_of Db * star_of (f x)" | |
| 214 | in add_commute [THEN subst]) | |
| 215 | apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym] | |
| 216 | Infinitesimal_add Infinitesimal_mult | |
| 217 | Infinitesimal_star_of_mult | |
| 218 | Infinitesimal_star_of_mult2 | |
| 219 | simp add: add_assoc [symmetric]) | |
| 220 | done | |
| 221 | ||
| 222 | text{*Multiplying by a constant*}
 | |
| 223 | lemma NSDERIV_cmult: "NSDERIV f x :> D | |
| 224 | ==> NSDERIV (%x. c * f x) x :> c*D" | |
| 225 | apply (simp only: times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff | |
| 226 | minus_mult_right right_diff_distrib [symmetric]) | |
| 227 | apply (erule NSLIM_const [THEN NSLIM_mult]) | |
| 228 | done | |
| 229 | ||
| 230 | text{*Negation of function*}
 | |
| 231 | lemma NSDERIV_minus: "NSDERIV f x :> D ==> NSDERIV (%x. -(f x)) x :> -D" | |
| 232 | proof (simp add: NSDERIV_NSLIM_iff) | |
| 233 | assume "(\<lambda>h. (f (x + h) - f x) / h) -- 0 --NS> D" | |
| 234 | hence deriv: "(\<lambda>h. - ((f(x+h) - f x) / h)) -- 0 --NS> - D" | |
| 235 | by (rule NSLIM_minus) | |
| 236 | have "\<forall>h. - ((f (x + h) - f x) / h) = (- f (x + h) + f x) / h" | |
| 237 | by (simp add: minus_divide_left diff_def) | |
| 238 | with deriv | |
| 239 | show "(\<lambda>h. (- f (x + h) + f x) / h) -- 0 --NS> - D" by simp | |
| 240 | qed | |
| 241 | ||
| 242 | text{*Subtraction*}
 | |
| 243 | lemma NSDERIV_add_minus: "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] ==> NSDERIV (%x. f x + -g x) x :> Da + -Db" | |
| 244 | by (blast dest: NSDERIV_add NSDERIV_minus) | |
| 245 | ||
| 246 | lemma NSDERIV_diff: | |
| 247 | "[| NSDERIV f x :> Da; NSDERIV g x :> Db |] | |
| 248 | ==> NSDERIV (%x. f x - g x) x :> Da-Db" | |
| 249 | apply (simp add: diff_minus) | |
| 250 | apply (blast intro: NSDERIV_add_minus) | |
| 251 | done | |
| 252 | ||
| 253 | text{*  Similarly to the above, the chain rule admits an entirely
 | |
| 254 | straightforward derivation. Compare this with Harrison's | |
| 255 | HOL proof of the chain rule, which proved to be trickier and | |
| 256 | required an alternative characterisation of differentiability- | |
| 257 | the so-called Carathedory derivative. Our main problem is | |
| 258 | manipulation of terms.*} | |
| 259 | ||
| 260 | (* lemmas *) | |
| 261 | ||
| 262 | lemma NSDERIV_zero: | |
| 263 | "[| NSDERIV g x :> D; | |
| 264 | ( *f* g) (star_of x + xa) = star_of (g x); | |
| 265 | xa \<in> Infinitesimal; | |
| 266 | xa \<noteq> 0 | |
| 267 | |] ==> D = 0" | |
| 268 | apply (simp add: nsderiv_def) | |
| 269 | apply (drule bspec, auto) | |
| 270 | done | |
| 271 | ||
| 272 | (* can be proved differently using NSLIM_isCont_iff *) | |
| 273 | lemma NSDERIV_approx: | |
| 274 | "[| NSDERIV f x :> D; h \<in> Infinitesimal; h \<noteq> 0 |] | |
| 275 | ==> ( *f* f) (star_of x + h) - star_of (f x) \<approx> 0" | |
| 276 | apply (simp add: nsderiv_def) | |
| 277 | apply (simp add: mem_infmal_iff [symmetric]) | |
| 278 | apply (rule Infinitesimal_ratio) | |
| 279 | apply (rule_tac [3] approx_star_of_HFinite, auto) | |
| 280 | done | |
| 281 | ||
| 282 | (*--------------------------------------------------------------- | |
| 283 | from one version of differentiability | |
| 284 | ||
| 285 | f(x) - f(a) | |
| 286 | --------------- \<approx> Db | |
| 287 | x - a | |
| 288 | ---------------------------------------------------------------*) | |
| 289 | ||
| 290 | lemma NSDERIVD1: "[| NSDERIV f (g x) :> Da; | |
| 291 | ( *f* g) (star_of(x) + xa) \<noteq> star_of (g x); | |
| 292 | ( *f* g) (star_of(x) + xa) \<approx> star_of (g x) | |
| 293 | |] ==> (( *f* f) (( *f* g) (star_of(x) + xa)) | |
| 294 | - star_of (f (g x))) | |
| 295 | / (( *f* g) (star_of(x) + xa) - star_of (g x)) | |
| 296 | \<approx> star_of(Da)" | |
| 297 | by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def diff_minus [symmetric]) | |
| 298 | ||
| 299 | (*-------------------------------------------------------------- | |
| 300 | from other version of differentiability | |
| 301 | ||
| 302 | f(x + h) - f(x) | |
| 303 | ----------------- \<approx> Db | |
| 304 | h | |
| 305 | --------------------------------------------------------------*) | |
| 306 | ||
| 307 | lemma NSDERIVD2: "[| NSDERIV g x :> Db; xa \<in> Infinitesimal; xa \<noteq> 0 |] | |
| 308 | ==> (( *f* g) (star_of(x) + xa) - star_of(g x)) / xa | |
| 309 | \<approx> star_of(Db)" | |
| 310 | by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel) | |
| 311 | ||
| 312 | lemma lemma_chain: "(z::'a::real_normed_field star) \<noteq> 0 ==> x*y = (x*inverse(z))*(z*y)" | |
| 313 | proof - | |
| 314 | assume z: "z \<noteq> 0" | |
| 315 | have "x * y = x * (inverse z * z) * y" by (simp add: z) | |
| 316 | thus ?thesis by (simp add: mult_assoc) | |
| 317 | qed | |
| 318 | ||
| 319 | text{*This proof uses both definitions of differentiability.*}
 | |
| 320 | lemma NSDERIV_chain: "[| NSDERIV f (g x) :> Da; NSDERIV g x :> Db |] | |
| 321 | ==> NSDERIV (f o g) x :> Da * Db" | |
| 322 | apply (simp (no_asm_simp) add: NSDERIV_NSLIM_iff NSLIM_def | |
| 323 | mem_infmal_iff [symmetric]) | |
| 324 | apply clarify | |
| 325 | apply (frule_tac f = g in NSDERIV_approx) | |
| 326 | apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric]) | |
| 327 | apply (case_tac "( *f* g) (star_of (x) + xa) = star_of (g x) ") | |
| 328 | apply (drule_tac g = g in NSDERIV_zero) | |
| 329 | apply (auto simp add: divide_inverse) | |
| 330 | apply (rule_tac z1 = "( *f* g) (star_of (x) + xa) - star_of (g x) " and y1 = "inverse xa" in lemma_chain [THEN ssubst]) | |
| 331 | apply (erule hypreal_not_eq_minus_iff [THEN iffD1]) | |
| 332 | apply (rule approx_mult_star_of) | |
| 333 | apply (simp_all add: divide_inverse [symmetric]) | |
| 334 | apply (blast intro: NSDERIVD1 approx_minus_iff [THEN iffD2]) | |
| 335 | apply (blast intro: NSDERIVD2) | |
| 336 | done | |
| 337 | ||
| 338 | text{*Differentiation of natural number powers*}
 | |
| 339 | lemma NSDERIV_Id [simp]: "NSDERIV (%x. x) x :> 1" | |
| 340 | by (simp add: NSDERIV_NSLIM_iff NSLIM_def del: divide_self_if) | |
| 341 | ||
| 342 | lemma NSDERIV_cmult_Id [simp]: "NSDERIV (op * c) x :> c" | |
| 343 | by (cut_tac c = c and x = x in NSDERIV_Id [THEN NSDERIV_cmult], simp) | |
| 344 | ||
| 345 | (*Can't get rid of x \<noteq> 0 because it isn't continuous at zero*) | |
| 346 | lemma NSDERIV_inverse: | |
| 31017 | 347 |   fixes x :: "'a::{real_normed_field}"
 | 
| 27468 | 348 | shows "x \<noteq> 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ Suc (Suc 0)))" | 
| 349 | apply (simp add: nsderiv_def) | |
| 350 | apply (rule ballI, simp, clarify) | |
| 351 | apply (frule (1) Infinitesimal_add_not_zero) | |
| 352 | apply (simp add: add_commute) | |
| 353 | (*apply (auto simp add: starfun_inverse_inverse realpow_two | |
| 354 | simp del: minus_mult_left [symmetric] minus_mult_right [symmetric])*) | |
| 355 | apply (simp add: inverse_add nonzero_inverse_mult_distrib [symmetric] power_Suc | |
| 356 | nonzero_inverse_minus_eq [symmetric] add_ac mult_ac diff_def | |
| 357 | del: inverse_mult_distrib inverse_minus_eq | |
| 358 | minus_mult_left [symmetric] minus_mult_right [symmetric]) | |
| 359 | apply (subst mult_commute, simp add: nonzero_mult_divide_cancel_right) | |
| 360 | apply (simp (no_asm_simp) add: mult_assoc [symmetric] left_distrib | |
| 361 | del: minus_mult_left [symmetric] minus_mult_right [symmetric]) | |
| 362 | apply (rule_tac y = "inverse (- (star_of x * star_of x))" in approx_trans) | |
| 363 | apply (rule inverse_add_Infinitesimal_approx2) | |
| 364 | apply (auto dest!: hypreal_of_real_HFinite_diff_Infinitesimal | |
| 365 | simp add: inverse_minus_eq [symmetric] HFinite_minus_iff) | |
| 366 | apply (rule Infinitesimal_HFinite_mult, auto) | |
| 367 | done | |
| 368 | ||
| 369 | subsubsection {* Equivalence of NS and Standard definitions *}
 | |
| 370 | ||
| 371 | lemma divideR_eq_divide: "x /\<^sub>R y = x / y" | |
| 372 | by (simp add: real_scaleR_def divide_inverse mult_commute) | |
| 373 | ||
| 374 | text{*Now equivalence between NSDERIV and DERIV*}
 | |
| 375 | lemma NSDERIV_DERIV_iff: "(NSDERIV f x :> D) = (DERIV f x :> D)" | |
| 376 | by (simp add: deriv_def NSDERIV_NSLIM_iff LIM_NSLIM_iff) | |
| 377 | ||
| 378 | (* NS version *) | |
| 379 | lemma NSDERIV_pow: "NSDERIV (%x. x ^ n) x :> real n * (x ^ (n - Suc 0))" | |
| 380 | by (simp add: NSDERIV_DERIV_iff DERIV_pow) | |
| 381 | ||
| 382 | text{*Derivative of inverse*}
 | |
| 383 | ||
| 384 | lemma NSDERIV_inverse_fun: | |
| 31017 | 385 |   fixes x :: "'a::{real_normed_field}"
 | 
| 27468 | 386 | shows "[| NSDERIV f x :> d; f(x) \<noteq> 0 |] | 
| 387 | ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ Suc (Suc 0))))" | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
28562diff
changeset | 388 | by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: power_Suc) | 
| 27468 | 389 | |
| 390 | text{*Derivative of quotient*}
 | |
| 391 | ||
| 392 | lemma NSDERIV_quotient: | |
| 31017 | 393 |   fixes x :: "'a::{real_normed_field}"
 | 
| 27468 | 394 | shows "[| NSDERIV f x :> d; NSDERIV g x :> e; g(x) \<noteq> 0 |] | 
| 395 | ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) | |
| 396 | - (e*f(x))) / (g(x) ^ Suc (Suc 0))" | |
| 30273 
ecd6f0ca62ea
declare power_Suc [simp]; remove redundant type-specific versions of power_Suc
 huffman parents: 
28562diff
changeset | 397 | by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: power_Suc) | 
| 27468 | 398 | |
| 399 | lemma CARAT_NSDERIV: "NSDERIV f x :> l ==> | |
| 400 | \<exists>g. (\<forall>z. f z - f x = g z * (z-x)) & isNSCont g x & g x = l" | |
| 401 | by (auto simp add: NSDERIV_DERIV_iff isNSCont_isCont_iff CARAT_DERIV | |
| 402 | mult_commute) | |
| 403 | ||
| 404 | lemma hypreal_eq_minus_iff3: "(x = y + z) = (x + -z = (y::hypreal))" | |
| 405 | by auto | |
| 406 | ||
| 407 | lemma CARAT_DERIVD: | |
| 408 | assumes all: "\<forall>z. f z - f x = g z * (z-x)" | |
| 409 | and nsc: "isNSCont g x" | |
| 410 | shows "NSDERIV f x :> g x" | |
| 411 | proof - | |
| 412 | from nsc | |
| 413 | have "\<forall>w. w \<noteq> star_of x \<and> w \<approx> star_of x \<longrightarrow> | |
| 414 | ( *f* g) w * (w - star_of x) / (w - star_of x) \<approx> | |
| 415 | star_of (g x)" | |
| 416 | by (simp add: isNSCont_def nonzero_mult_divide_cancel_right) | |
| 417 | thus ?thesis using all | |
| 418 | by (simp add: NSDERIV_iff2 starfun_if_eq cong: if_cong) | |
| 419 | qed | |
| 420 | ||
| 421 | subsubsection {* Differentiability predicate *}
 | |
| 422 | ||
| 423 | lemma NSdifferentiableD: "f NSdifferentiable x ==> \<exists>D. NSDERIV f x :> D" | |
| 424 | by (simp add: NSdifferentiable_def) | |
| 425 | ||
| 426 | lemma NSdifferentiableI: "NSDERIV f x :> D ==> f NSdifferentiable x" | |
| 427 | by (force simp add: NSdifferentiable_def) | |
| 428 | ||
| 429 | ||
| 430 | subsection {*(NS) Increment*}
 | |
| 431 | lemma incrementI: | |
| 432 | "f NSdifferentiable x ==> | |
| 433 | increment f x h = ( *f* f) (hypreal_of_real(x) + h) - | |
| 434 | hypreal_of_real (f x)" | |
| 435 | by (simp add: increment_def) | |
| 436 | ||
| 437 | lemma incrementI2: "NSDERIV f x :> D ==> | |
| 438 | increment f x h = ( *f* f) (hypreal_of_real(x) + h) - | |
| 439 | hypreal_of_real (f x)" | |
| 440 | apply (erule NSdifferentiableI [THEN incrementI]) | |
| 441 | done | |
| 442 | ||
| 443 | (* The Increment theorem -- Keisler p. 65 *) | |
| 444 | lemma increment_thm: "[| NSDERIV f x :> D; h \<in> Infinitesimal; h \<noteq> 0 |] | |
| 445 | ==> \<exists>e \<in> Infinitesimal. increment f x h = hypreal_of_real(D)*h + e*h" | |
| 446 | apply (frule_tac h = h in incrementI2, simp add: nsderiv_def) | |
| 447 | apply (drule bspec, auto) | |
| 448 | apply (drule bex_Infinitesimal_iff2 [THEN iffD2], clarify) | |
| 449 | apply (frule_tac b1 = "hypreal_of_real (D) + y" | |
| 450 | in hypreal_mult_right_cancel [THEN iffD2]) | |
| 451 | apply (erule_tac [2] V = "(( *f* f) (hypreal_of_real (x) + h) - hypreal_of_real (f x)) / h = hypreal_of_real (D) + y" in thin_rl) | |
| 452 | apply assumption | |
| 453 | apply (simp add: times_divide_eq_right [symmetric]) | |
| 454 | apply (auto simp add: left_distrib) | |
| 455 | done | |
| 456 | ||
| 457 | lemma increment_thm2: | |
| 458 | "[| NSDERIV f x :> D; h \<approx> 0; h \<noteq> 0 |] | |
| 459 | ==> \<exists>e \<in> Infinitesimal. increment f x h = | |
| 460 | hypreal_of_real(D)*h + e*h" | |
| 461 | by (blast dest!: mem_infmal_iff [THEN iffD2] intro!: increment_thm) | |
| 462 | ||
| 463 | ||
| 464 | lemma increment_approx_zero: "[| NSDERIV f x :> D; h \<approx> 0; h \<noteq> 0 |] | |
| 465 | ==> increment f x h \<approx> 0" | |
| 466 | apply (drule increment_thm2, | |
| 467 | auto intro!: Infinitesimal_HFinite_mult2 HFinite_add simp add: left_distrib [symmetric] mem_infmal_iff [symmetric]) | |
| 468 | apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) | |
| 469 | done | |
| 470 | ||
| 471 | end |