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