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