| author | paulson <lp15@cam.ac.uk> | 
| Mon, 09 Dec 2019 16:37:26 +0000 | |
| changeset 71260 | 308baf6b450a | 
| parent 70723 | 4e39d87c9737 | 
| child 80914 | d97fdabd9e2b | 
| permissions | -rw-r--r-- | 
| 62479 | 1 | (* Title: HOL/Nonstandard_Analysis/HDeriv.thy | 
| 2 | Author: Jacques D. Fleuriot | |
| 3 | Copyright: 1998 University of Cambridge | |
| 27468 | 4 | Conversion to Isar and new proofs by Lawrence C Paulson, 2004 | 
| 5 | *) | |
| 6 | ||
| 64435 | 7 | section \<open>Differentiation (Nonstandard)\<close> | 
| 27468 | 8 | |
| 9 | theory HDeriv | |
| 64435 | 10 | imports HLim | 
| 27468 | 11 | begin | 
| 12 | ||
| 64435 | 13 | text \<open>Nonstandard Definitions.\<close> | 
| 27468 | 14 | |
| 64435 | 15 | definition nsderiv :: "['a::real_normed_field \<Rightarrow> 'a, 'a, 'a] \<Rightarrow> bool" | 
| 16 |     ("(NSDERIV (_)/ (_)/ :> (_))" [1000, 1000, 60] 60)
 | |
| 17 | where "NSDERIV f x :> D \<longleftrightarrow> | |
| 18 |     (\<forall>h \<in> Infinitesimal - {0}. (( *f* f)(star_of x + h) - star_of (f x)) / h \<approx> star_of D)"
 | |
| 27468 | 19 | |
| 64435 | 20 | definition NSdifferentiable :: "['a::real_normed_field \<Rightarrow> 'a, 'a] \<Rightarrow> bool" | 
| 21 | (infixl "NSdifferentiable" 60) | |
| 22 | where "f NSdifferentiable x \<longleftrightarrow> (\<exists>D. NSDERIV f x :> D)" | |
| 27468 | 23 | |
| 64435 | 24 | definition increment :: "(real \<Rightarrow> real) \<Rightarrow> real \<Rightarrow> hypreal \<Rightarrow> hypreal" | 
| 25 | where "increment f x h = | |
| 26 | (SOME inc. f NSdifferentiable x \<and> inc = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x))" | |
| 27468 | 27 | |
| 28 | ||
| 61975 | 29 | subsection \<open>Derivatives\<close> | 
| 27468 | 30 | |
| 64435 | 31 | lemma DERIV_NS_iff: "(DERIV f x :> D) \<longleftrightarrow> (\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D" | 
| 32 | by (simp add: DERIV_def LIM_NSLIM_iff) | |
| 27468 | 33 | |
| 64435 | 34 | lemma NS_DERIV_D: "DERIV f x :> D \<Longrightarrow> (\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D" | 
| 35 | by (simp add: DERIV_def LIM_NSLIM_iff) | |
| 27468 | 36 | |
| 37 | lemma Infinitesimal_of_hypreal: | |
| 64435 | 38 | "x \<in> Infinitesimal \<Longrightarrow> (( *f* of_real) x::'a::real_normed_div_algebra star) \<in> Infinitesimal" | 
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 39 | by (metis Infinitesimal_of_hypreal_iff of_hypreal_def) | 
| 27468 | 40 | |
| 64435 | 41 | lemma of_hypreal_eq_0_iff: "\<And>x. (( *f* of_real) x = (0::'a::real_algebra_1 star)) = (x = 0)" | 
| 42 | by transfer (rule of_real_eq_0_iff) | |
| 27468 | 43 | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 44 | lemma NSDeriv_unique: | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 45 | assumes "NSDERIV f x :> D" "NSDERIV f x :> E" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 46 | shows "NSDERIV f x :> D \<Longrightarrow> NSDERIV f x :> E \<Longrightarrow> D = E" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 47 | proof - | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 48 |   have "\<exists>s. (s::'a star) \<in> Infinitesimal - {0}"
 | 
| 70723 
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
 paulson <lp15@cam.ac.uk> parents: 
70217diff
changeset | 49 | by (metis Diff_iff HDeriv.of_hypreal_eq_0_iff Infinitesimal_epsilon Infinitesimal_of_hypreal epsilon_not_zero singletonD) | 
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 50 | with assms show ?thesis | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 51 | by (meson approx_trans3 nsderiv_def star_of_approx_iff) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 52 | qed | 
| 27468 | 53 | |
| 64435 | 54 | text \<open>First \<open>NSDERIV\<close> in terms of \<open>NSLIM\<close>.\<close> | 
| 27468 | 55 | |
| 64435 | 56 | text \<open>First equivalence.\<close> | 
| 57 | lemma NSDERIV_NSLIM_iff: "(NSDERIV f x :> D) \<longleftrightarrow> (\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D" | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 58 | by (auto simp add: nsderiv_def NSLIM_def starfun_lambda_cancel mem_infmal_iff) | 
| 27468 | 59 | |
| 64435 | 60 | text \<open>Second equivalence.\<close> | 
| 61 | lemma NSDERIV_NSLIM_iff2: "(NSDERIV f x :> D) \<longleftrightarrow> (\<lambda>z. (f z - f x) / (z - x)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S D" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53755diff
changeset | 62 | by (simp add: NSDERIV_NSLIM_iff DERIV_LIM_iff LIM_NSLIM_iff [symmetric]) | 
| 27468 | 63 | |
| 64435 | 64 | text \<open>While we're at it!\<close> | 
| 27468 | 65 | lemma NSDERIV_iff2: | 
| 64435 | 66 | "(NSDERIV f x :> D) \<longleftrightarrow> | 
| 64604 | 67 | (\<forall>w. w \<noteq> star_of x \<and> w \<approx> star_of x \<longrightarrow> ( *f* (\<lambda>z. (f z - f x) / (z - x))) w \<approx> star_of D)" | 
| 64435 | 68 | by (simp add: NSDERIV_NSLIM_iff2 NSLIM_def) | 
| 27468 | 69 | |
| 70 | lemma NSDERIVD5: | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 71 | "\<lbrakk>NSDERIV f x :> D; u \<approx> hypreal_of_real x\<rbrakk> \<Longrightarrow> | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 72 | ( *f* (\<lambda>z. f z - f x)) u \<approx> hypreal_of_real D * (u - hypreal_of_real x)" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 73 | unfolding NSDERIV_iff2 | 
| 64435 | 74 | apply (case_tac "u = hypreal_of_real x", auto) | 
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 75 | by (metis (mono_tags, lifting) HFinite_star_of Infinitesimal_ratio approx_def approx_minus_iff approx_mult_subst approx_star_of_HFinite approx_sym mult_zero_right right_minus_eq) | 
| 27468 | 76 | |
| 77 | lemma NSDERIVD4: | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 78 | "\<lbrakk>NSDERIV f x :> D; h \<in> Infinitesimal\<rbrakk> | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 79 | \<Longrightarrow> ( *f* f)(hypreal_of_real x + h) - hypreal_of_real (f x) \<approx> hypreal_of_real D * h" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 80 | apply (clarsimp simp add: nsderiv_def) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 81 | apply (case_tac "h = 0", simp) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 82 | by (meson DiffI Infinitesimal_approx Infinitesimal_ratio Infinitesimal_star_of_mult2 approx_star_of_HFinite singletonD) | 
| 27468 | 83 | |
| 64435 | 84 | text \<open>Differentiability implies continuity nice and simple "algebraic" proof.\<close> | 
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 85 | lemma NSDERIV_isNSCont: | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 86 | assumes "NSDERIV f x :> D" shows "isNSCont f x" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 87 | unfolding isNSCont_NSLIM_iff NSLIM_def | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 88 | proof clarify | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 89 | fix x' | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 90 | assume "x' \<noteq> star_of x" "x' \<approx> star_of x" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 91 |   then have m0: "x' - star_of x \<in> Infinitesimal - {0}"
 | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 92 | using bex_Infinitesimal_iff by auto | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 93 | then have "(( *f* f) x' - star_of (f x)) / (x' - star_of x) \<approx> star_of D" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 94 | by (metis \<open>x' \<approx> star_of x\<close> add_diff_cancel_left' assms bex_Infinitesimal_iff2 nsderiv_def) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 95 | then have "(( *f* f) x' - star_of (f x)) / (x' - star_of x) \<in> HFinite" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 96 | by (metis approx_star_of_HFinite) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 97 | then show "( *f* f) x' \<approx> star_of (f x)" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 98 | by (metis (no_types) Diff_iff Infinitesimal_ratio m0 bex_Infinitesimal_iff insert_iff) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 99 | qed | 
| 27468 | 100 | |
| 64435 | 101 | text \<open>Differentiation rules for combinations of functions | 
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 102 | follow from clear, straightforward, algebraic manipulations.\<close> | 
| 64435 | 103 | |
| 104 | text \<open>Constant function.\<close> | |
| 27468 | 105 | |
| 106 | (* use simple constant nslimit theorem *) | |
| 64435 | 107 | lemma NSDERIV_const [simp]: "NSDERIV (\<lambda>x. k) x :> 0" | 
| 108 | by (simp add: NSDERIV_NSLIM_iff) | |
| 27468 | 109 | |
| 64435 | 110 | text \<open>Sum of functions- proved easily.\<close> | 
| 27468 | 111 | |
| 64435 | 112 | lemma NSDERIV_add: | 
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 113 | assumes "NSDERIV f x :> Da" "NSDERIV g x :> Db" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 114 | shows "NSDERIV (\<lambda>x. f x + g x) x :> Da + Db" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 115 | proof - | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 116 | have "((\<lambda>x. f x + g x) has_field_derivative Da + Db) (at x)" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 117 | using assms DERIV_NS_iff NSDERIV_NSLIM_iff field_differentiable_add by blast | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 118 | then show ?thesis | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 119 | by (simp add: DERIV_NS_iff NSDERIV_NSLIM_iff) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 120 | qed | 
| 64435 | 121 | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 122 | text \<open>Product of functions - Proof is simple.\<close> | 
| 27468 | 123 | |
| 64435 | 124 | lemma NSDERIV_mult: | 
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 125 | assumes "NSDERIV g x :> Db" "NSDERIV f x :> Da" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 126 | shows "NSDERIV (\<lambda>x. f x * g x) x :> (Da * g x) + (Db * f x)" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 127 | proof - | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 128 | have "(f has_field_derivative Da) (at x)" "(g has_field_derivative Db) (at x)" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 129 | using assms by (simp_all add: DERIV_NS_iff NSDERIV_NSLIM_iff) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 130 | then have "((\<lambda>a. f a * g a) has_field_derivative Da * g x + Db * f x) (at x)" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 131 | using DERIV_mult by blast | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 132 | then show ?thesis | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 133 | by (simp add: DERIV_NS_iff NSDERIV_NSLIM_iff) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 134 | qed | 
| 27468 | 135 | |
| 64435 | 136 | text \<open>Multiplying by a constant.\<close> | 
| 137 | lemma NSDERIV_cmult: "NSDERIV f x :> D \<Longrightarrow> NSDERIV (\<lambda>x. c * f x) x :> c * D" | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 138 | unfolding times_divide_eq_right [symmetric] NSDERIV_NSLIM_iff | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 139 | minus_mult_right right_diff_distrib [symmetric] | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 140 | by (erule NSLIM_const [THEN NSLIM_mult]) | 
| 64435 | 141 | |
| 142 | text \<open>Negation of function.\<close> | |
| 143 | lemma NSDERIV_minus: "NSDERIV f x :> D \<Longrightarrow> NSDERIV (\<lambda>x. - f x) x :> - D" | |
| 27468 | 144 | proof (simp add: NSDERIV_NSLIM_iff) | 
| 61971 | 145 | assume "(\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D" | 
| 64435 | 146 | then have deriv: "(\<lambda>h. - ((f(x+h) - f x) / h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S - D" | 
| 27468 | 147 | by (rule NSLIM_minus) | 
| 148 | have "\<forall>h. - ((f (x + h) - f x) / h) = (- f (x + h) + f x) / h" | |
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53755diff
changeset | 149 | by (simp add: minus_divide_left) | 
| 64435 | 150 | with deriv have "(\<lambda>h. (- f (x + h) + f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S - D" | 
| 151 | by simp | |
| 152 | then show "(\<lambda>h. (f (x + h) - f x) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S D \<Longrightarrow> (\<lambda>h. (f x - f (x + h)) / h) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S - D" | |
| 153 | by simp | |
| 27468 | 154 | qed | 
| 155 | ||
| 64435 | 156 | text \<open>Subtraction.\<close> | 
| 157 | lemma NSDERIV_add_minus: | |
| 158 | "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (\<lambda>x. f x + - g x) x :> Da + - Db" | |
| 159 | by (blast dest: NSDERIV_add NSDERIV_minus) | |
| 27468 | 160 | |
| 161 | lemma NSDERIV_diff: | |
| 64435 | 162 | "NSDERIV f x :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (\<lambda>x. f x - g x) x :> Da - Db" | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53755diff
changeset | 163 | using NSDERIV_add_minus [of f x Da g Db] by simp | 
| 27468 | 164 | |
| 64435 | 165 | text \<open>Similarly to the above, the chain rule admits an entirely | 
| 166 | straightforward derivation. Compare this with Harrison's | |
| 167 | HOL proof of the chain rule, which proved to be trickier and | |
| 168 | required an alternative characterisation of differentiability- | |
| 169 | the so-called Carathedory derivative. Our main problem is | |
| 170 | manipulation of terms.\<close> | |
| 27468 | 171 | |
| 64435 | 172 | |
| 173 | subsection \<open>Lemmas\<close> | |
| 27468 | 174 | |
| 175 | lemma NSDERIV_zero: | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 176 | "\<lbrakk>NSDERIV g x :> D; ( *f* g) (star_of x + y) = star_of (g x); y \<in> Infinitesimal; y \<noteq> 0\<rbrakk> | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 177 | \<Longrightarrow> D = 0" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 178 | by (force simp add: nsderiv_def) | 
| 27468 | 179 | |
| 64435 | 180 | text \<open>Can be proved differently using \<open>NSLIM_isCont_iff\<close>.\<close> | 
| 27468 | 181 | lemma NSDERIV_approx: | 
| 64435 | 182 | "NSDERIV f x :> D \<Longrightarrow> h \<in> Infinitesimal \<Longrightarrow> h \<noteq> 0 \<Longrightarrow> | 
| 183 | ( *f* f) (star_of x + h) - star_of (f x) \<approx> 0" | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 184 | by (meson DiffI Infinitesimal_ratio approx_star_of_HFinite mem_infmal_iff nsderiv_def singletonD) | 
| 27468 | 185 | |
| 64435 | 186 | text \<open>From one version of differentiability | 
| 27468 | 187 | |
| 64435 | 188 | \<open>f x - f a\<close> | 
| 189 | \<open>-------------- \<approx> Db\<close> | |
| 190 | \<open>x - a\<close> | |
| 191 | \<close> | |
| 27468 | 192 | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 193 | lemma NSDERIVD1: | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 194 | "\<lbrakk>NSDERIV f (g x) :> Da; | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 195 | ( *f* g) (star_of x + y) \<noteq> star_of (g x); | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 196 | ( *f* g) (star_of x + y) \<approx> star_of (g x)\<rbrakk> | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 197 | \<Longrightarrow> (( *f* f) (( *f* g) (star_of x + y)) - | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 198 | star_of (f (g x))) / (( *f* g) (star_of x + y) - star_of (g x)) \<approx> | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 199 | star_of Da" | 
| 64435 | 200 | by (auto simp add: NSDERIV_NSLIM_iff2 NSLIM_def) | 
| 27468 | 201 | |
| 64435 | 202 | text \<open>From other version of differentiability | 
| 27468 | 203 | |
| 64435 | 204 | \<open>f (x + h) - f x\<close> | 
| 205 | \<open>------------------ \<approx> Db\<close> | |
| 206 | \<open>h\<close> | |
| 207 | \<close> | |
| 27468 | 208 | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 209 | lemma NSDERIVD2: "[| NSDERIV g x :> Db; y \<in> Infinitesimal; y \<noteq> 0 |] | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 210 | ==> (( *f* g) (star_of(x) + y) - star_of(g x)) / y | 
| 27468 | 211 | \<approx> star_of(Db)" | 
| 64435 | 212 | by (auto simp add: NSDERIV_NSLIM_iff NSLIM_def mem_infmal_iff starfun_lambda_cancel) | 
| 27468 | 213 | |
| 64435 | 214 | text \<open>This proof uses both definitions of differentiability.\<close> | 
| 215 | lemma NSDERIV_chain: | |
| 216 | "NSDERIV f (g x) :> Da \<Longrightarrow> NSDERIV g x :> Db \<Longrightarrow> NSDERIV (f \<circ> g) x :> Da * Db" | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 217 | using DERIV_NS_iff DERIV_chain NSDERIV_NSLIM_iff by blast | 
| 27468 | 218 | |
| 64435 | 219 | text \<open>Differentiation of natural number powers.\<close> | 
| 220 | lemma NSDERIV_Id [simp]: "NSDERIV (\<lambda>x. x) x :> 1" | |
| 221 | by (simp add: NSDERIV_NSLIM_iff NSLIM_def del: divide_self_if) | |
| 27468 | 222 | |
| 69064 
5840724b1d71
Prefix form of infix with * on either side no longer needs special treatment
 nipkow parents: 
68644diff
changeset | 223 | lemma NSDERIV_cmult_Id [simp]: "NSDERIV ((*) c) x :> c" | 
| 64435 | 224 | using NSDERIV_Id [THEN NSDERIV_cmult] by simp | 
| 27468 | 225 | |
| 226 | lemma NSDERIV_inverse: | |
| 53755 | 227 | fixes x :: "'a::real_normed_field" | 
| 69597 | 228 | assumes "x \<noteq> 0" \<comment> \<open>can't get rid of \<^term>\<open>x \<noteq> 0\<close> because it isn't continuous at zero\<close> | 
| 53755 | 229 | shows "NSDERIV (\<lambda>x. inverse x) x :> - (inverse x ^ Suc (Suc 0))" | 
| 230 | proof - | |
| 64435 | 231 |   {
 | 
| 232 | fix h :: "'a star" | |
| 53755 | 233 | assume h_Inf: "h \<in> Infinitesimal" | 
| 64435 | 234 | from this assms have not_0: "star_of x + h \<noteq> 0" | 
| 235 | by (rule Infinitesimal_add_not_zero) | |
| 53755 | 236 | assume "h \<noteq> 0" | 
| 64435 | 237 | from h_Inf have "h * star_of x \<in> Infinitesimal" | 
| 238 | by (rule Infinitesimal_HFinite_mult) simp | |
| 53755 | 239 | with assms have "inverse (- (h * star_of x) + - (star_of x * star_of x)) \<approx> | 
| 240 | inverse (- (star_of x * star_of x))" | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 241 | proof - | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 242 | have "- (h * star_of x) + - (star_of x * star_of x) \<approx> - (star_of x * star_of x)" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 243 | using \<open>h * star_of x \<in> Infinitesimal\<close> assms bex_Infinitesimal_iff by fastforce | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 244 | then show ?thesis | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 245 | by (metis assms mult_eq_0_iff neg_equal_0_iff_equal star_of_approx_inverse star_of_minus star_of_mult) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 246 | qed | 
| 61975 | 247 | moreover from not_0 \<open>h \<noteq> 0\<close> assms | 
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 248 | have "inverse (- (h * star_of x) + - (star_of x * star_of x)) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 249 | = (inverse (star_of x + h) - inverse (star_of x)) / h" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 250 | by (simp add: division_ring_inverse_diff inverse_mult_distrib [symmetric] | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 251 | inverse_minus_eq [symmetric] algebra_simps) | 
| 54230 
b1d955791529
more simplification rules on unary and binary minus
 haftmann parents: 
53755diff
changeset | 252 | ultimately have "(inverse (star_of x + h) - inverse (star_of x)) / h \<approx> | 
| 53755 | 253 | - (inverse (star_of x) * inverse (star_of x))" | 
| 64435 | 254 | using assms by simp | 
| 255 | } | |
| 256 | then show ?thesis by (simp add: nsderiv_def) | |
| 53755 | 257 | qed | 
| 27468 | 258 | |
| 64435 | 259 | |
| 61975 | 260 | subsubsection \<open>Equivalence of NS and Standard definitions\<close> | 
| 27468 | 261 | |
| 262 | lemma divideR_eq_divide: "x /\<^sub>R y = x / y" | |
| 64435 | 263 | by (simp add: divide_inverse mult.commute) | 
| 27468 | 264 | |
| 64435 | 265 | text \<open>Now equivalence between \<open>NSDERIV\<close> and \<open>DERIV\<close>.\<close> | 
| 266 | lemma NSDERIV_DERIV_iff: "NSDERIV f x :> D \<longleftrightarrow> DERIV f x :> D" | |
| 267 | by (simp add: DERIV_def NSDERIV_NSLIM_iff LIM_NSLIM_iff) | |
| 27468 | 268 | |
| 64435 | 269 | text \<open>NS version.\<close> | 
| 270 | lemma NSDERIV_pow: "NSDERIV (\<lambda>x. x ^ n) x :> real n * (x ^ (n - Suc 0))" | |
| 271 | by (simp add: NSDERIV_DERIV_iff DERIV_pow) | |
| 27468 | 272 | |
| 64435 | 273 | text \<open>Derivative of inverse.\<close> | 
| 27468 | 274 | lemma NSDERIV_inverse_fun: | 
| 64435 | 275 | "NSDERIV f x :> d \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> | 
| 276 | NSDERIV (\<lambda>x. inverse (f x)) x :> (- (d * inverse (f x ^ Suc (Suc 0))))" | |
| 277 |   for x :: "'a::{real_normed_field}"
 | |
| 278 | by (simp add: NSDERIV_DERIV_iff DERIV_inverse_fun del: power_Suc) | |
| 27468 | 279 | |
| 64435 | 280 | text \<open>Derivative of quotient.\<close> | 
| 27468 | 281 | lemma NSDERIV_quotient: | 
| 64435 | 282 | fixes x :: "'a::real_normed_field" | 
| 283 | shows "NSDERIV f x :> d \<Longrightarrow> NSDERIV g x :> e \<Longrightarrow> g x \<noteq> 0 \<Longrightarrow> | |
| 284 | NSDERIV (\<lambda>y. f y / g y) x :> (d * g x - (e * f x)) / (g x ^ Suc (Suc 0))" | |
| 285 | by (simp add: NSDERIV_DERIV_iff DERIV_quotient del: power_Suc) | |
| 27468 | 286 | |
| 64435 | 287 | lemma CARAT_NSDERIV: | 
| 288 | "NSDERIV f x :> l \<Longrightarrow> \<exists>g. (\<forall>z. f z - f x = g z * (z - x)) \<and> isNSCont g x \<and> g x = l" | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 289 | by (simp add: CARAT_DERIV NSDERIV_DERIV_iff isNSCont_isCont_iff) | 
| 27468 | 290 | |
| 64435 | 291 | lemma hypreal_eq_minus_iff3: "x = y + z \<longleftrightarrow> x + - z = y" | 
| 292 | for x y z :: hypreal | |
| 293 | by auto | |
| 27468 | 294 | |
| 295 | lemma CARAT_DERIVD: | |
| 64435 | 296 | assumes all: "\<forall>z. f z - f x = g z * (z - x)" | 
| 297 | and nsc: "isNSCont g x" | |
| 27468 | 298 | shows "NSDERIV f x :> g x" | 
| 299 | proof - | |
| 64435 | 300 | from nsc have "\<forall>w. w \<noteq> star_of x \<and> w \<approx> star_of x \<longrightarrow> | 
| 301 | ( *f* g) w * (w - star_of x) / (w - star_of x) \<approx> star_of (g x)" | |
| 302 | by (simp add: isNSCont_def) | |
| 303 | with all show ?thesis | |
| 27468 | 304 | by (simp add: NSDERIV_iff2 starfun_if_eq cong: if_cong) | 
| 305 | qed | |
| 306 | ||
| 64435 | 307 | |
| 61975 | 308 | subsubsection \<open>Differentiability predicate\<close> | 
| 27468 | 309 | |
| 64435 | 310 | lemma NSdifferentiableD: "f NSdifferentiable x \<Longrightarrow> \<exists>D. NSDERIV f x :> D" | 
| 311 | by (simp add: NSdifferentiable_def) | |
| 27468 | 312 | |
| 64435 | 313 | lemma NSdifferentiableI: "NSDERIV f x :> D \<Longrightarrow> f NSdifferentiable x" | 
| 314 | by (force simp add: NSdifferentiable_def) | |
| 27468 | 315 | |
| 316 | ||
| 61975 | 317 | subsection \<open>(NS) Increment\<close> | 
| 27468 | 318 | |
| 64435 | 319 | lemma incrementI: | 
| 320 | "f NSdifferentiable x \<Longrightarrow> | |
| 321 | increment f x h = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)" | |
| 322 | by (simp add: increment_def) | |
| 323 | ||
| 324 | lemma incrementI2: | |
| 325 | "NSDERIV f x :> D \<Longrightarrow> | |
| 326 | increment f x h = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)" | |
| 327 | by (erule NSdifferentiableI [THEN incrementI]) | |
| 27468 | 328 | |
| 64435 | 329 | text \<open>The Increment theorem -- Keisler p. 65.\<close> | 
| 330 | lemma increment_thm: | |
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 331 | assumes "NSDERIV f x :> D" "h \<in> Infinitesimal" "h \<noteq> 0" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 332 | shows "\<exists>e \<in> Infinitesimal. increment f x h = hypreal_of_real D * h + e * h" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 333 | proof - | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 334 | have inc: "increment f x h = ( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 335 | using assms(1) incrementI2 by auto | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 336 | have "(( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)) / h \<approx> hypreal_of_real D" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 337 | by (simp add: NSDERIVD2 assms) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 338 | then obtain y where "y \<in> Infinitesimal" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 339 | "(( *f* f) (hypreal_of_real x + h) - hypreal_of_real (f x)) / h = hypreal_of_real D + y" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 340 | by (metis bex_Infinitesimal_iff2) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 341 | then have "increment f x h / h = hypreal_of_real D + y" | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 342 | by (metis inc) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 343 | then show ?thesis | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 344 | by (metis (no_types) \<open>y \<in> Infinitesimal\<close> \<open>h \<noteq> 0\<close> distrib_right mult.commute nonzero_mult_div_cancel_left times_divide_eq_right) | 
| 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 345 | qed | 
| 27468 | 346 | |
| 64435 | 347 | lemma increment_approx_zero: "NSDERIV f x :> D \<Longrightarrow> h \<approx> 0 \<Longrightarrow> h \<noteq> 0 \<Longrightarrow> increment f x h \<approx> 0" | 
| 68644 
242d298526a3
de-applying and simplifying proofs
 paulson <lp15@cam.ac.uk> parents: 
67399diff
changeset | 348 | by (simp add: NSDERIV_approx incrementI2 mem_infmal_iff) | 
| 27468 | 349 | |
| 350 | end |