author | wenzelm |
Fri, 17 Nov 2006 02:20:03 +0100 | |
changeset 21404 | eb85850d3eb7 |
parent 20752 | 09cf0e407a45 |
child 21792 | 266a1056a2a3 |
permissions | -rw-r--r-- |
13957 | 1 |
(* Title : CLim.thy |
2 |
Author : Jacques D. Fleuriot |
|
3 |
Copyright : 2001 University of Edinburgh |
|
14469 | 4 |
Conversion to Isar and new proofs by Lawrence C Paulson, 2004 |
13957 | 5 |
*) |
6 |
||
14469 | 7 |
header{*Limits, Continuity and Differentiation for Complex Functions*} |
8 |
||
15131 | 9 |
theory CLim |
15140 | 10 |
imports CSeries |
15131 | 11 |
begin |
14405 | 12 |
|
13 |
(*not in simpset?*) |
|
14 |
declare hypreal_epsilon_not_zero [simp] |
|
15 |
||
16 |
(*??generalize*) |
|
17 |
lemma lemma_complex_mult_inverse_squared [simp]: |
|
18 |
"x \<noteq> (0::complex) \<Longrightarrow> (x * inverse(x) ^ 2) = inverse x" |
|
14469 | 19 |
by (simp add: numeral_2_eq_2) |
14405 | 20 |
|
21 |
text{*Changing the quantified variable. Install earlier?*} |
|
14738 | 22 |
lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) = (\<forall>x. P (x-a))"; |
14405 | 23 |
apply auto |
24 |
apply (drule_tac x="x+a" in spec) |
|
25 |
apply (simp add: diff_minus add_assoc) |
|
26 |
done |
|
27 |
||
28 |
lemma complex_add_minus_iff [simp]: "(x + - a = (0::complex)) = (x=a)" |
|
29 |
by (simp add: diff_eq_eq diff_minus [symmetric]) |
|
30 |
||
31 |
lemma complex_add_eq_0_iff [iff]: "(x+y = (0::complex)) = (y = -x)" |
|
32 |
apply auto |
|
33 |
apply (drule sym [THEN diff_eq_eq [THEN iffD2]], auto) |
|
34 |
done |
|
13957 | 35 |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
36 |
abbreviation |
14405 | 37 |
CLIM :: "[complex=>complex,complex,complex] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
38 |
("((_)/ -- (_)/ --C> (_))" [60, 0, 60] 60) where |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
39 |
"CLIM == LIM" |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
40 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
41 |
abbreviation |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
42 |
NSCLIM :: "[complex=>complex,complex,complex] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
43 |
("((_)/ -- (_)/ --NSC> (_))" [60, 0, 60] 60) where |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
44 |
"NSCLIM == NSLIM" |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
45 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
46 |
abbreviation |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
47 |
(* f: C --> R *) |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
48 |
CRLIM :: "[complex=>real,complex,real] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
49 |
("((_)/ -- (_)/ --CR> (_))" [60, 0, 60] 60) where |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
50 |
"CRLIM == LIM" |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
51 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
52 |
abbreviation |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
53 |
NSCRLIM :: "[complex=>real,complex,real] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
54 |
("((_)/ -- (_)/ --NSCR> (_))" [60, 0, 60] 60) where |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
55 |
"NSCRLIM == NSLIM" |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
56 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
57 |
abbreviation |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
58 |
isContc :: "[complex=>complex,complex] => bool" where |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
59 |
"isContc == isCont" |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
60 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
61 |
abbreviation |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
62 |
(* NS definition dispenses with limit notions *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
63 |
isNSContc :: "[complex=>complex,complex] => bool" where |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
64 |
"isNSContc == isNSCont" |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
65 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
66 |
abbreviation |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
67 |
isContCR :: "[complex=>real,complex] => bool" where |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
68 |
"isContCR == isCont" |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
69 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
70 |
abbreviation |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
71 |
(* NS definition dispenses with limit notions *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
72 |
isNSContCR :: "[complex=>real,complex] => bool" where |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
73 |
"isNSContCR == isNSCont" |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
74 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
75 |
abbreviation |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
76 |
isUContc :: "(complex=>complex) => bool" where |
20752 | 77 |
"isUContc == isUCont" |
78 |
||
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
79 |
abbreviation |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
80 |
isNSUContc :: "(complex=>complex) => bool" where |
20752 | 81 |
"isNSUContc == isNSUCont" |
82 |
||
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
83 |
|
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
84 |
lemma CLIM_def: |
19765 | 85 |
"f -- a --C> L = |
86 |
(\<forall>r. 0 < r --> |
|
14405 | 87 |
(\<exists>s. 0 < s & (\<forall>x. (x \<noteq> a & (cmod(x - a) < s) |
19765 | 88 |
--> cmod(f x - L) < r))))" |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
89 |
by (rule LIM_def) |
13957 | 90 |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
91 |
lemma NSCLIM_def: |
19765 | 92 |
"f -- a --NSC> L = (\<forall>x. (x \<noteq> hcomplex_of_complex a & |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
93 |
x @= hcomplex_of_complex a |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
94 |
--> ( *f* f) x @= hcomplex_of_complex L))" |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
95 |
by (rule NSLIM_def) |
13957 | 96 |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
97 |
lemma CRLIM_def: |
19765 | 98 |
"f -- a --CR> L = |
99 |
(\<forall>r. 0 < r --> |
|
14405 | 100 |
(\<exists>s. 0 < s & (\<forall>x. (x \<noteq> a & (cmod(x - a) < s) |
19765 | 101 |
--> abs(f x - L) < r))))" |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
102 |
by (simp add: LIM_def) |
13957 | 103 |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
104 |
lemma NSCRLIM_def: |
19765 | 105 |
"f -- a --NSCR> L = (\<forall>x. (x \<noteq> hcomplex_of_complex a & |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
106 |
x @= hcomplex_of_complex a |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
107 |
--> ( *f* f) x @= hypreal_of_real L))" |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
108 |
by (rule NSLIM_def) |
13957 | 109 |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
110 |
lemma isContc_def: |
19765 | 111 |
"isContc f a = (f -- a --C> (f a))" |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
112 |
by (rule isCont_def) |
13957 | 113 |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
114 |
lemma isNSContc_def: |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
115 |
"isNSContc f a = (\<forall>y. y @= hcomplex_of_complex a --> |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
116 |
( *f* f) y @= hcomplex_of_complex (f a))" |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
117 |
by (rule isNSCont_def) |
13957 | 118 |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
119 |
lemma isContCR_def: |
19765 | 120 |
"isContCR f a = (f -- a --CR> (f a))" |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
121 |
by (rule isCont_def) |
13957 | 122 |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
123 |
lemma isNSContCR_def: |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
124 |
"isNSContCR f a = (\<forall>y. y @= hcomplex_of_complex a --> |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
125 |
( *f* f) y @= hypreal_of_real (f a))" |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
126 |
by (rule isNSCont_def) |
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
127 |
|
20752 | 128 |
lemma isUContc_def: |
129 |
"isUContc f = (\<forall>r. 0 < r --> |
|
130 |
(\<exists>s. 0 < s & (\<forall>x y. cmod(x - y) < s |
|
131 |
--> cmod(f x - f y) < r)))" |
|
132 |
by (rule isUCont_def) |
|
133 |
||
134 |
lemma isNSUContc_def: |
|
135 |
"isNSUContc f = (\<forall>x y. x @= y --> ( *f* f) x @= ( *f* f) y)" |
|
136 |
by (rule isNSUCont_def) |
|
137 |
||
138 |
||
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
139 |
(* differentiation: D is derivative of function f at x *) |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
140 |
definition |
14405 | 141 |
cderiv:: "[complex=>complex,complex,complex] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
142 |
("(CDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) where |
19765 | 143 |
"CDERIV f x :> D = ((%h. (f(x + h) - f(x))/h) -- 0 --C> D)" |
13957 | 144 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
145 |
definition |
14405 | 146 |
nscderiv :: "[complex=>complex,complex,complex] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
147 |
("(NSCDERIV (_)/ (_)/ :> (_))" [60, 0, 60] 60) where |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
148 |
"NSCDERIV f x :> D = (\<forall>h \<in> Infinitesimal - {0}. |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
149 |
(( *f* f)(hcomplex_of_complex x + h) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
150 |
- hcomplex_of_complex (f x))/h @= hcomplex_of_complex D)" |
13957 | 151 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
152 |
definition |
14405 | 153 |
cdifferentiable :: "[complex=>complex,complex] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
154 |
(infixl "cdifferentiable" 60) where |
19765 | 155 |
"f cdifferentiable x = (\<exists>D. CDERIV f x :> D)" |
14405 | 156 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
157 |
definition |
14405 | 158 |
NSCdifferentiable :: "[complex=>complex,complex] => bool" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20752
diff
changeset
|
159 |
(infixl "NSCdifferentiable" 60) where |
19765 | 160 |
"f NSCdifferentiable x = (\<exists>D. NSCDERIV f x :> D)" |
14405 | 161 |
|
162 |
||
163 |
subsection{*Limit of Complex to Complex Function*} |
|
164 |
||
165 |
lemma NSCLIM_NSCRLIM_Re: "f -- a --NSC> L ==> (%x. Re(f x)) -- a --NSCR> Re(L)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
166 |
by (simp add: NSLIM_def starfunC_approx_Re_Im_iff |
14469 | 167 |
hRe_hcomplex_of_complex) |
168 |
||
14405 | 169 |
|
14469 | 170 |
lemma NSCLIM_NSCRLIM_Im: "f -- a --NSC> L ==> (%x. Im(f x)) -- a --NSCR> Im(L)" |
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
171 |
by (simp add: NSLIM_def starfunC_approx_Re_Im_iff |
14469 | 172 |
hIm_hcomplex_of_complex) |
14405 | 173 |
|
174 |
lemma CLIM_NSCLIM: |
|
175 |
"f -- x --C> L ==> f -- x --NSC> L" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
176 |
by (rule LIM_NSLIM) |
14405 | 177 |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
178 |
lemma eq_Abs_star_ALL: "(\<forall>t. P t) = (\<forall>X. P (star_n X))" |
14405 | 179 |
apply auto |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
180 |
apply (rule_tac x = t in star_cases, auto) |
14405 | 181 |
done |
182 |
||
183 |
lemma lemma_CLIM: |
|
184 |
"\<forall>s. 0 < s --> (\<exists>xa. xa \<noteq> x & |
|
185 |
cmod (xa - x) < s & r \<le> cmod (f xa - L)) |
|
186 |
==> \<forall>(n::nat). \<exists>xa. xa \<noteq> x & |
|
187 |
cmod(xa - x) < inverse(real(Suc n)) & r \<le> cmod(f xa - L)" |
|
188 |
apply clarify |
|
189 |
apply (cut_tac n1 = n in real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive], auto) |
|
190 |
done |
|
191 |
||
192 |
||
193 |
lemma lemma_skolemize_CLIM2: |
|
194 |
"\<forall>s. 0 < s --> (\<exists>xa. xa \<noteq> x & |
|
195 |
cmod (xa - x) < s & r \<le> cmod (f xa - L)) |
|
196 |
==> \<exists>X. \<forall>(n::nat). X n \<noteq> x & |
|
197 |
cmod(X n - x) < inverse(real(Suc n)) & r \<le> cmod(f (X n) - L)" |
|
198 |
apply (drule lemma_CLIM) |
|
199 |
apply (drule choice, blast) |
|
200 |
done |
|
201 |
||
202 |
lemma lemma_csimp: |
|
203 |
"\<forall>n. X n \<noteq> x & |
|
204 |
cmod (X n - x) < inverse (real(Suc n)) & |
|
205 |
r \<le> cmod (f (X n) - L) ==> |
|
206 |
\<forall>n. cmod (X n - x) < inverse (real(Suc n))" |
|
207 |
by auto |
|
208 |
||
209 |
lemma NSCLIM_CLIM: |
|
210 |
"f -- x --NSC> L ==> f -- x --C> L" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
211 |
by (rule NSLIM_LIM) |
14405 | 212 |
|
213 |
||
214 |
text{*First key result*} |
|
215 |
theorem CLIM_NSCLIM_iff: "(f -- x --C> L) = (f -- x --NSC> L)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
216 |
by (rule LIM_NSLIM_iff) |
14405 | 217 |
|
218 |
||
219 |
subsection{*Limit of Complex to Real Function*} |
|
220 |
||
221 |
lemma CRLIM_NSCRLIM: "f -- x --CR> L ==> f -- x --NSCR> L" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
222 |
by (rule LIM_NSLIM) |
14405 | 223 |
|
224 |
lemma lemma_CRLIM: |
|
225 |
"\<forall>s. 0 < s --> (\<exists>xa. xa \<noteq> x & |
|
226 |
cmod (xa - x) < s & r \<le> abs (f xa - L)) |
|
227 |
==> \<forall>(n::nat). \<exists>xa. xa \<noteq> x & |
|
228 |
cmod(xa - x) < inverse(real(Suc n)) & r \<le> abs (f xa - L)" |
|
229 |
apply clarify |
|
230 |
apply (cut_tac n1 = n in real_of_nat_Suc_gt_zero [THEN positive_imp_inverse_positive], auto) |
|
231 |
done |
|
232 |
||
233 |
lemma lemma_skolemize_CRLIM2: |
|
234 |
"\<forall>s. 0 < s --> (\<exists>xa. xa \<noteq> x & |
|
235 |
cmod (xa - x) < s & r \<le> abs (f xa - L)) |
|
236 |
==> \<exists>X. \<forall>(n::nat). X n \<noteq> x & |
|
237 |
cmod(X n - x) < inverse(real(Suc n)) & r \<le> abs (f (X n) - L)" |
|
238 |
apply (drule lemma_CRLIM) |
|
239 |
apply (drule choice, blast) |
|
240 |
done |
|
241 |
||
242 |
lemma lemma_crsimp: |
|
243 |
"\<forall>n. X n \<noteq> x & |
|
244 |
cmod (X n - x) < inverse (real(Suc n)) & |
|
245 |
r \<le> abs (f (X n) - L) ==> |
|
246 |
\<forall>n. cmod (X n - x) < inverse (real(Suc n))" |
|
247 |
by auto |
|
248 |
||
249 |
lemma NSCRLIM_CRLIM: "f -- x --NSCR> L ==> f -- x --CR> L" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
250 |
by (rule NSLIM_LIM) |
14405 | 251 |
|
252 |
text{*Second key result*} |
|
253 |
theorem CRLIM_NSCRLIM_iff: "(f -- x --CR> L) = (f -- x --NSCR> L)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
254 |
by (rule LIM_NSLIM_iff) |
14405 | 255 |
|
256 |
(** get this result easily now **) |
|
257 |
lemma CLIM_CRLIM_Re: "f -- a --C> L ==> (%x. Re(f x)) -- a --CR> Re(L)" |
|
258 |
by (auto dest: NSCLIM_NSCRLIM_Re simp add: CLIM_NSCLIM_iff CRLIM_NSCRLIM_iff [symmetric]) |
|
259 |
||
260 |
lemma CLIM_CRLIM_Im: "f -- a --C> L ==> (%x. Im(f x)) -- a --CR> Im(L)" |
|
261 |
by (auto dest: NSCLIM_NSCRLIM_Im simp add: CLIM_NSCLIM_iff CRLIM_NSCRLIM_iff [symmetric]) |
|
262 |
||
263 |
lemma CLIM_cnj: "f -- a --C> L ==> (%x. cnj (f x)) -- a --C> cnj L" |
|
14469 | 264 |
by (simp add: CLIM_def complex_cnj_diff [symmetric]) |
14405 | 265 |
|
266 |
lemma CLIM_cnj_iff: "((%x. cnj (f x)) -- a --C> cnj L) = (f -- a --C> L)" |
|
14469 | 267 |
by (simp add: CLIM_def complex_cnj_diff [symmetric]) |
14405 | 268 |
|
269 |
(*** NSLIM_add hence CLIM_add *) |
|
270 |
||
271 |
lemma NSCLIM_add: |
|
272 |
"[| f -- x --NSC> l; g -- x --NSC> m |] |
|
273 |
==> (%x. f(x) + g(x)) -- x --NSC> (l + m)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
274 |
by (rule NSLIM_add) |
14405 | 275 |
|
276 |
lemma CLIM_add: |
|
277 |
"[| f -- x --C> l; g -- x --C> m |] |
|
278 |
==> (%x. f(x) + g(x)) -- x --C> (l + m)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
279 |
by (rule LIM_add) |
14405 | 280 |
|
281 |
(*** NSLIM_mult hence CLIM_mult *) |
|
282 |
||
283 |
lemma NSCLIM_mult: |
|
284 |
"[| f -- x --NSC> l; g -- x --NSC> m |] |
|
285 |
==> (%x. f(x) * g(x)) -- x --NSC> (l * m)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
286 |
by (rule NSLIM_mult) |
14405 | 287 |
|
288 |
lemma CLIM_mult: |
|
289 |
"[| f -- x --C> l; g -- x --C> m |] |
|
290 |
==> (%x. f(x) * g(x)) -- x --C> (l * m)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
291 |
by (rule LIM_mult2) |
14405 | 292 |
|
293 |
(*** NSCLIM_const and CLIM_const ***) |
|
294 |
||
295 |
lemma NSCLIM_const [simp]: "(%x. k) -- x --NSC> k" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
296 |
by (rule NSLIM_const) |
14405 | 297 |
|
298 |
lemma CLIM_const [simp]: "(%x. k) -- x --C> k" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
299 |
by (rule LIM_const) |
14405 | 300 |
|
301 |
(*** NSCLIM_minus and CLIM_minus ***) |
|
302 |
||
303 |
lemma NSCLIM_minus: "f -- a --NSC> L ==> (%x. -f(x)) -- a --NSC> -L" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
304 |
by (rule NSLIM_minus) |
14405 | 305 |
|
306 |
lemma CLIM_minus: "f -- a --C> L ==> (%x. -f(x)) -- a --C> -L" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
307 |
by (rule LIM_minus) |
14405 | 308 |
|
309 |
(*** NSCLIM_diff hence CLIM_diff ***) |
|
310 |
||
311 |
lemma NSCLIM_diff: |
|
312 |
"[| f -- x --NSC> l; g -- x --NSC> m |] |
|
313 |
==> (%x. f(x) - g(x)) -- x --NSC> (l - m)" |
|
14469 | 314 |
by (simp add: diff_minus NSCLIM_add NSCLIM_minus) |
14405 | 315 |
|
316 |
lemma CLIM_diff: |
|
317 |
"[| f -- x --C> l; g -- x --C> m |] |
|
318 |
==> (%x. f(x) - g(x)) -- x --C> (l - m)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
319 |
by (rule LIM_diff) |
14405 | 320 |
|
321 |
(*** NSCLIM_inverse and hence CLIM_inverse *) |
|
322 |
||
323 |
lemma NSCLIM_inverse: |
|
324 |
"[| f -- a --NSC> L; L \<noteq> 0 |] |
|
325 |
==> (%x. inverse(f(x))) -- a --NSC> (inverse L)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
326 |
by (rule NSLIM_inverse) |
14405 | 327 |
|
328 |
lemma CLIM_inverse: |
|
329 |
"[| f -- a --C> L; L \<noteq> 0 |] |
|
330 |
==> (%x. inverse(f(x))) -- a --C> (inverse L)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
331 |
by (rule LIM_inverse) |
14405 | 332 |
|
333 |
(*** NSCLIM_zero, CLIM_zero, etc. ***) |
|
334 |
||
335 |
lemma NSCLIM_zero: "f -- a --NSC> l ==> (%x. f(x) - l) -- a --NSC> 0" |
|
14469 | 336 |
apply (simp add: diff_minus) |
14405 | 337 |
apply (rule_tac a1 = l in right_minus [THEN subst]) |
14469 | 338 |
apply (rule NSCLIM_add, auto) |
14405 | 339 |
done |
340 |
||
341 |
lemma CLIM_zero: "f -- a --C> l ==> (%x. f(x) - l) -- a --C> 0" |
|
342 |
by (simp add: CLIM_NSCLIM_iff NSCLIM_zero) |
|
343 |
||
344 |
lemma NSCLIM_zero_cancel: "(%x. f(x) - l) -- x --NSC> 0 ==> f -- x --NSC> l" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
345 |
by (rule NSLIM_zero_cancel) |
14405 | 346 |
|
347 |
lemma CLIM_zero_cancel: "(%x. f(x) - l) -- x --C> 0 ==> f -- x --C> l" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
348 |
by (rule LIM_zero_cancel) |
14405 | 349 |
|
350 |
(*** NSCLIM_not zero and hence CLIM_not_zero ***) |
|
13957 | 351 |
|
14405 | 352 |
lemma NSCLIM_not_zero: "k \<noteq> 0 ==> ~ ((%x. k) -- x --NSC> 0)" |
17373
27509e72f29e
removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents:
17318
diff
changeset
|
353 |
apply (auto simp del: star_of_zero simp add: NSCLIM_def) |
14405 | 354 |
apply (rule_tac x = "hcomplex_of_complex x + hcomplex_of_hypreal epsilon" in exI) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
355 |
apply (auto intro: Infinitesimal_add_approx_self [THEN approx_sym] |
17373
27509e72f29e
removed duplicated lemmas; convert more proofs to transfer principle
huffman
parents:
17318
diff
changeset
|
356 |
simp del: star_of_zero) |
14405 | 357 |
done |
358 |
||
359 |
(* [| k \<noteq> 0; (%x. k) -- x --NSC> 0 |] ==> R *) |
|
360 |
lemmas NSCLIM_not_zeroE = NSCLIM_not_zero [THEN notE, standard] |
|
361 |
||
362 |
lemma CLIM_not_zero: "k \<noteq> 0 ==> ~ ((%x. k) -- x --C> 0)" |
|
363 |
by (simp add: CLIM_NSCLIM_iff NSCLIM_not_zero) |
|
364 |
||
365 |
(*** NSCLIM_const hence CLIM_const ***) |
|
366 |
||
367 |
lemma NSCLIM_const_eq: "(%x. k) -- x --NSC> L ==> k = L" |
|
368 |
apply (rule ccontr) |
|
369 |
apply (drule NSCLIM_zero) |
|
370 |
apply (rule NSCLIM_not_zeroE [of "k-L"], auto) |
|
371 |
done |
|
372 |
||
373 |
lemma CLIM_const_eq: "(%x. k) -- x --C> L ==> k = L" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
374 |
by (rule LIM_const_eq) |
14405 | 375 |
|
376 |
(*** NSCLIM and hence CLIM are unique ***) |
|
377 |
||
378 |
lemma NSCLIM_unique: "[| f -- x --NSC> L; f -- x --NSC> M |] ==> L = M" |
|
379 |
apply (drule NSCLIM_minus) |
|
380 |
apply (drule NSCLIM_add, assumption) |
|
381 |
apply (auto dest!: NSCLIM_const_eq [symmetric]) |
|
382 |
done |
|
383 |
||
384 |
lemma CLIM_unique: "[| f -- x --C> L; f -- x --C> M |] ==> L = M" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
385 |
by (rule LIM_unique) |
14405 | 386 |
|
387 |
(*** NSCLIM_mult_zero and CLIM_mult_zero ***) |
|
388 |
||
389 |
lemma NSCLIM_mult_zero: |
|
390 |
"[| f -- x --NSC> 0; g -- x --NSC> 0 |] ==> (%x. f(x)*g(x)) -- x --NSC> 0" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
391 |
by (rule NSLIM_mult_zero) |
14405 | 392 |
|
393 |
lemma CLIM_mult_zero: |
|
394 |
"[| f -- x --C> 0; g -- x --C> 0 |] ==> (%x. f(x)*g(x)) -- x --C> 0" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
395 |
by (rule LIM_mult_zero) |
14405 | 396 |
|
397 |
(*** NSCLIM_self hence CLIM_self ***) |
|
398 |
||
399 |
lemma NSCLIM_self: "(%x. x) -- a --NSC> a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
400 |
by (rule NSLIM_self) |
14405 | 401 |
|
402 |
lemma CLIM_self: "(%x. x) -- a --C> a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
403 |
by (rule LIM_self) |
14405 | 404 |
|
405 |
(** another equivalence result **) |
|
406 |
lemma NSCLIM_NSCRLIM_iff: |
|
407 |
"(f -- x --NSC> L) = ((%y. cmod(f y - L)) -- x --NSCR> 0)" |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
408 |
apply (auto simp add: NSCLIM_def NSCRLIM_def Infinitesimal_approx_minus [symmetric] Infinitesimal_hcmod_iff) |
14405 | 409 |
apply (auto dest!: spec) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
410 |
apply (rule_tac [!] x = xa in star_cases) |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
411 |
apply (auto simp add: star_n_diff starfun hcmod mem_infmal_iff star_of_def) |
14405 | 412 |
done |
413 |
||
414 |
(** much, much easier standard proof **) |
|
415 |
lemma CLIM_CRLIM_iff: "(f -- x --C> L) = ((%y. cmod(f y - L)) -- x --CR> 0)" |
|
416 |
by (simp add: CLIM_def CRLIM_def) |
|
417 |
||
418 |
(* so this is nicer nonstandard proof *) |
|
419 |
lemma NSCLIM_NSCRLIM_iff2: |
|
420 |
"(f -- x --NSC> L) = ((%y. cmod(f y - L)) -- x --NSCR> 0)" |
|
14469 | 421 |
by (simp add: CRLIM_NSCRLIM_iff [symmetric] CLIM_CRLIM_iff CLIM_NSCLIM_iff [symmetric]) |
14405 | 422 |
|
423 |
lemma NSCLIM_NSCRLIM_Re_Im_iff: |
|
424 |
"(f -- a --NSC> L) = ((%x. Re(f x)) -- a --NSCR> Re(L) & |
|
425 |
(%x. Im(f x)) -- a --NSCR> Im(L))" |
|
426 |
apply (auto intro: NSCLIM_NSCRLIM_Re NSCLIM_NSCRLIM_Im) |
|
427 |
apply (auto simp add: NSCLIM_def NSCRLIM_def) |
|
428 |
apply (auto dest!: spec) |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
429 |
apply (rule_tac x = x in star_cases) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
430 |
apply (simp add: approx_approx_iff starfun star_of_def) |
14405 | 431 |
done |
432 |
||
433 |
lemma CLIM_CRLIM_Re_Im_iff: |
|
434 |
"(f -- a --C> L) = ((%x. Re(f x)) -- a --CR> Re(L) & |
|
435 |
(%x. Im(f x)) -- a --CR> Im(L))" |
|
436 |
by (simp add: CLIM_NSCLIM_iff CRLIM_NSCRLIM_iff NSCLIM_NSCRLIM_Re_Im_iff) |
|
437 |
||
438 |
||
439 |
subsection{*Continuity*} |
|
440 |
||
441 |
lemma isNSContcD: |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
442 |
"[| isNSContc f a; y @= hcomplex_of_complex a |] |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
443 |
==> ( *f* f) y @= hcomplex_of_complex (f a)" |
14405 | 444 |
by (simp add: isNSContc_def) |
445 |
||
446 |
lemma isNSContc_NSCLIM: "isNSContc f a ==> f -- a --NSC> (f a) " |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
447 |
by (rule isNSCont_NSLIM) |
14405 | 448 |
|
449 |
lemma NSCLIM_isNSContc: |
|
450 |
"f -- a --NSC> (f a) ==> isNSContc f a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
451 |
by (rule NSLIM_isNSCont) |
14405 | 452 |
|
453 |
text{*Nonstandard continuity can be defined using NS Limit in |
|
454 |
similar fashion to standard definition of continuity*} |
|
455 |
||
456 |
lemma isNSContc_NSCLIM_iff: "(isNSContc f a) = (f -- a --NSC> (f a))" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
457 |
by (rule isNSCont_NSLIM_iff) |
14405 | 458 |
|
459 |
lemma isNSContc_CLIM_iff: "(isNSContc f a) = (f -- a --C> (f a))" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
460 |
by (rule isNSCont_LIM_iff) |
14405 | 461 |
|
462 |
(*** key result for continuity ***) |
|
463 |
lemma isNSContc_isContc_iff: "(isNSContc f a) = (isContc f a)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
464 |
by (rule isNSCont_isCont_iff) |
14405 | 465 |
|
466 |
lemma isContc_isNSContc: "isContc f a ==> isNSContc f a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
467 |
by (rule isCont_isNSCont) |
14405 | 468 |
|
469 |
lemma isNSContc_isContc: "isNSContc f a ==> isContc f a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
470 |
by (rule isNSCont_isCont) |
14405 | 471 |
|
472 |
||
473 |
text{*Alternative definition of continuity*} |
|
474 |
lemma NSCLIM_h_iff: "(f -- a --NSC> L) = ((%h. f(a + h)) -- 0 --NSC> L)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
475 |
by (rule NSLIM_h_iff) |
14405 | 476 |
|
477 |
lemma NSCLIM_isContc_iff: |
|
478 |
"(f -- a --NSC> f a) = ((%h. f(a + h)) -- 0 --NSC> f a)" |
|
479 |
by (rule NSCLIM_h_iff) |
|
480 |
||
481 |
lemma CLIM_isContc_iff: "(f -- a --C> f a) = ((%h. f(a + h)) -- 0 --C> f(a))" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
482 |
by (rule LIM_isCont_iff) |
14405 | 483 |
|
484 |
lemma isContc_iff: "(isContc f x) = ((%h. f(x + h)) -- 0 --C> f(x))" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
485 |
by (rule isCont_iff) |
14405 | 486 |
|
487 |
lemma isContc_add: |
|
488 |
"[| isContc f a; isContc g a |] ==> isContc (%x. f(x) + g(x)) a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
489 |
by (rule isCont_add) |
14405 | 490 |
|
491 |
lemma isContc_mult: |
|
492 |
"[| isContc f a; isContc g a |] ==> isContc (%x. f(x) * g(x)) a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
493 |
by (rule isCont_mult) |
14405 | 494 |
|
495 |
||
496 |
lemma isContc_o: "[| isContc f a; isContc g (f a) |] ==> isContc (g o f) a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
497 |
by (rule isCont_o) |
14405 | 498 |
|
499 |
lemma isContc_o2: |
|
500 |
"[| isContc f a; isContc g (f a) |] ==> isContc (%x. g (f x)) a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
501 |
by (rule isCont_o2) |
14405 | 502 |
|
503 |
lemma isNSContc_minus: "isNSContc f a ==> isNSContc (%x. - f x) a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
504 |
by (rule isNSCont_minus) |
14405 | 505 |
|
506 |
lemma isContc_minus: "isContc f a ==> isContc (%x. - f x) a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
507 |
by (rule isCont_minus) |
14405 | 508 |
|
509 |
lemma isContc_inverse: |
|
510 |
"[| isContc f x; f x \<noteq> 0 |] ==> isContc (%x. inverse (f x)) x" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
511 |
by (rule isCont_inverse) |
14405 | 512 |
|
513 |
lemma isNSContc_inverse: |
|
514 |
"[| isNSContc f x; f x \<noteq> 0 |] ==> isNSContc (%x. inverse (f x)) x" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
515 |
by (rule isNSCont_inverse) |
14405 | 516 |
|
517 |
lemma isContc_diff: |
|
518 |
"[| isContc f a; isContc g a |] ==> isContc (%x. f(x) - g(x)) a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
519 |
by (rule isCont_diff) |
14405 | 520 |
|
521 |
lemma isContc_const [simp]: "isContc (%x. k) a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
522 |
by (rule isCont_const) |
14405 | 523 |
|
524 |
lemma isNSContc_const [simp]: "isNSContc (%x. k) a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
525 |
by (rule isNSCont_const) |
14405 | 526 |
|
527 |
||
528 |
subsection{*Functions from Complex to Reals*} |
|
529 |
||
530 |
lemma isNSContCRD: |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
531 |
"[| isNSContCR f a; y @= hcomplex_of_complex a |] |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
532 |
==> ( *f* f) y @= hypreal_of_real (f a)" |
14405 | 533 |
by (simp add: isNSContCR_def) |
534 |
||
535 |
lemma isNSContCR_NSCRLIM: "isNSContCR f a ==> f -- a --NSCR> (f a) " |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
536 |
by (rule isNSCont_NSLIM) |
14405 | 537 |
|
538 |
lemma NSCRLIM_isNSContCR: "f -- a --NSCR> (f a) ==> isNSContCR f a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
539 |
by (rule NSLIM_isNSCont) |
14405 | 540 |
|
541 |
lemma isNSContCR_NSCRLIM_iff: "(isNSContCR f a) = (f -- a --NSCR> (f a))" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
542 |
by (rule isNSCont_NSLIM_iff) |
14405 | 543 |
|
544 |
lemma isNSContCR_CRLIM_iff: "(isNSContCR f a) = (f -- a --CR> (f a))" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
545 |
by (rule isNSCont_LIM_iff) |
14405 | 546 |
|
547 |
(*** another key result for continuity ***) |
|
548 |
lemma isNSContCR_isContCR_iff: "(isNSContCR f a) = (isContCR f a)" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
549 |
by (rule isNSCont_isCont_iff) |
14405 | 550 |
|
551 |
lemma isContCR_isNSContCR: "isContCR f a ==> isNSContCR f a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
552 |
by (rule isCont_isNSCont) |
14405 | 553 |
|
554 |
lemma isNSContCR_isContCR: "isNSContCR f a ==> isContCR f a" |
|
20659
66b8455090b8
changed constants into abbreviations; shortened proofs
huffman
parents:
20563
diff
changeset
|
555 |
by (rule isNSCont_isCont) |
14405 | 556 |
|
557 |
lemma isNSContCR_cmod [simp]: "isNSContCR cmod (a)" |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
558 |
by (auto intro: approx_hcmod_approx |
14405 | 559 |
simp add: starfunCR_cmod hcmod_hcomplex_of_complex [symmetric] |
560 |
isNSContCR_def) |
|
561 |
||
562 |
lemma isContCR_cmod [simp]: "isContCR cmod (a)" |
|
14469 | 563 |
by (simp add: isNSContCR_isContCR_iff [symmetric]) |
14405 | 564 |
|
565 |
lemma isContc_isContCR_Re: "isContc f a ==> isContCR (%x. Re (f x)) a" |
|
566 |
by (simp add: isContc_def isContCR_def CLIM_CRLIM_Re) |
|
567 |
||
568 |
lemma isContc_isContCR_Im: "isContc f a ==> isContCR (%x. Im (f x)) a" |
|
569 |
by (simp add: isContc_def isContCR_def CLIM_CRLIM_Im) |
|
570 |
||
571 |
||
572 |
subsection{*Derivatives*} |
|
573 |
||
574 |
lemma CDERIV_iff: "(CDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --C> D)" |
|
575 |
by (simp add: cderiv_def) |
|
576 |
||
577 |
lemma CDERIV_NSC_iff: |
|
578 |
"(CDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NSC> D)" |
|
579 |
by (simp add: cderiv_def CLIM_NSCLIM_iff) |
|
580 |
||
581 |
lemma CDERIVD: "CDERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --C> D" |
|
582 |
by (simp add: cderiv_def) |
|
583 |
||
584 |
lemma NSC_DERIVD: "CDERIV f x :> D ==> (%h. (f(x + h) - f(x))/h) -- 0 --NSC> D" |
|
585 |
by (simp add: cderiv_def CLIM_NSCLIM_iff) |
|
586 |
||
587 |
text{*Uniqueness*} |
|
588 |
lemma CDERIV_unique: "[| CDERIV f x :> D; CDERIV f x :> E |] ==> D = E" |
|
589 |
by (simp add: cderiv_def CLIM_unique) |
|
590 |
||
591 |
(*** uniqueness: a nonstandard proof ***) |
|
592 |
lemma NSCDeriv_unique: "[| NSCDERIV f x :> D; NSCDERIV f x :> E |] ==> D = E" |
|
593 |
apply (simp add: nscderiv_def) |
|
594 |
apply (auto dest!: bspec [where x = "hcomplex_of_hypreal epsilon"] |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
595 |
intro!: inj_hcomplex_of_complex [THEN injD] dest: approx_trans3) |
14405 | 596 |
done |
13957 | 597 |
|
598 |
||
14405 | 599 |
subsection{* Differentiability*} |
600 |
||
601 |
lemma CDERIV_CLIM_iff: |
|
602 |
"((%h. (f(a + h) - f(a))/h) -- 0 --C> D) = |
|
603 |
((%x. (f(x) - f(a)) / (x - a)) -- a --C> D)" |
|
604 |
apply (simp add: CLIM_def) |
|
605 |
apply (rule_tac f=All in arg_cong) |
|
606 |
apply (rule ext) |
|
607 |
apply (rule imp_cong) |
|
608 |
apply (rule refl) |
|
609 |
apply (rule_tac f=Ex in arg_cong) |
|
610 |
apply (rule ext) |
|
611 |
apply (rule conj_cong) |
|
612 |
apply (rule refl) |
|
613 |
apply (rule trans) |
|
614 |
apply (rule all_shift [where a=a], simp) |
|
615 |
done |
|
616 |
||
617 |
lemma CDERIV_iff2: |
|
618 |
"(CDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z - x)) -- x --C> D)" |
|
619 |
by (simp add: cderiv_def CDERIV_CLIM_iff) |
|
620 |
||
621 |
||
622 |
subsection{* Equivalence of NS and Standard Differentiation*} |
|
623 |
||
624 |
(*** first equivalence ***) |
|
625 |
lemma NSCDERIV_NSCLIM_iff: |
|
626 |
"(NSCDERIV f x :> D) = ((%h. (f(x + h) - f(x))/h) -- 0 --NSC> D)" |
|
627 |
apply (simp add: nscderiv_def NSCLIM_def, auto) |
|
628 |
apply (drule_tac x = xa in bspec) |
|
629 |
apply (rule_tac [3] ccontr) |
|
630 |
apply (drule_tac [3] x = h in spec) |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
631 |
apply (auto simp add: mem_infmal_iff starfun_lambda_cancel) |
14405 | 632 |
done |
633 |
||
634 |
(*** 2nd equivalence ***) |
|
635 |
lemma NSCDERIV_NSCLIM_iff2: |
|
636 |
"(NSCDERIV f x :> D) = ((%z. (f(z) - f(x)) / (z - x)) -- x --NSC> D)" |
|
637 |
by (simp add: NSCDERIV_NSCLIM_iff CDERIV_CLIM_iff CLIM_NSCLIM_iff [symmetric]) |
|
638 |
||
639 |
lemma NSCDERIV_iff2: |
|
640 |
"(NSCDERIV f x :> D) = |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
641 |
(\<forall>xa. xa \<noteq> hcomplex_of_complex x & xa @= hcomplex_of_complex x --> |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
642 |
( *f* (%z. (f z - f x) / (z - x))) xa @= hcomplex_of_complex D)" |
14405 | 643 |
by (simp add: NSCDERIV_NSCLIM_iff2 NSCLIM_def) |
644 |
||
645 |
lemma NSCDERIV_CDERIV_iff: "(NSCDERIV f x :> D) = (CDERIV f x :> D)" |
|
646 |
by (simp add: cderiv_def NSCDERIV_NSCLIM_iff CLIM_NSCLIM_iff) |
|
647 |
||
648 |
lemma NSCDERIV_isNSContc: "NSCDERIV f x :> D ==> isNSContc f x" |
|
20563 | 649 |
apply (auto simp add: nscderiv_def isNSContc_NSCLIM_iff NSCLIM_def) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
650 |
apply (drule approx_minus_iff [THEN iffD1]) |
20563 | 651 |
apply (subgoal_tac "xa - (hcomplex_of_complex x) \<noteq> 0") |
14405 | 652 |
prefer 2 apply (simp add: compare_rls) |
20563 | 653 |
apply (drule_tac x = "xa - hcomplex_of_complex x" in bspec) |
654 |
apply (simp add: mem_infmal_iff) |
|
655 |
apply (simp add: add_assoc [symmetric]) |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
656 |
apply (auto simp add: mem_infmal_iff [symmetric] add_commute) |
20563 | 657 |
apply (drule_tac c = "xa - hcomplex_of_complex x" in approx_mult1) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
658 |
apply (auto intro: Infinitesimal_subset_HFinite [THEN subsetD] |
14405 | 659 |
simp add: mult_assoc) |
660 |
apply (drule_tac x3 = D in |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
661 |
HFinite_hcomplex_of_complex [THEN [2] Infinitesimal_HFinite_mult, |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
662 |
THEN mem_infmal_iff [THEN iffD1]]) |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
663 |
apply (blast intro: approx_trans mult_commute [THEN subst] approx_minus_iff [THEN iffD2]) |
14405 | 664 |
done |
665 |
||
666 |
lemma CDERIV_isContc: "CDERIV f x :> D ==> isContc f x" |
|
667 |
by (simp add: NSCDERIV_CDERIV_iff [symmetric] isNSContc_isContc_iff [symmetric] NSCDERIV_isNSContc) |
|
668 |
||
669 |
text{* Differentiation rules for combinations of functions follow by clear, |
|
670 |
straightforard algebraic manipulations*} |
|
671 |
||
672 |
(* use simple constant nslimit theorem *) |
|
673 |
lemma NSCDERIV_const [simp]: "(NSCDERIV (%x. k) x :> 0)" |
|
674 |
by (simp add: NSCDERIV_NSCLIM_iff) |
|
675 |
||
676 |
lemma CDERIV_const [simp]: "(CDERIV (%x. k) x :> 0)" |
|
677 |
by (simp add: NSCDERIV_CDERIV_iff [symmetric]) |
|
678 |
||
679 |
lemma NSCDERIV_add: |
|
680 |
"[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] |
|
681 |
==> NSCDERIV (%x. f x + g x) x :> Da + Db" |
|
682 |
apply (simp add: NSCDERIV_NSCLIM_iff NSCLIM_def, clarify) |
|
683 |
apply (auto dest!: spec simp add: add_divide_distrib diff_minus) |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
684 |
apply (drule_tac b = "hcomplex_of_complex Da" and d = "hcomplex_of_complex Db" in approx_add) |
14405 | 685 |
apply (auto simp add: add_ac) |
686 |
done |
|
687 |
||
688 |
lemma CDERIV_add: |
|
689 |
"[| CDERIV f x :> Da; CDERIV g x :> Db |] |
|
690 |
==> CDERIV (%x. f x + g x) x :> Da + Db" |
|
691 |
by (simp add: NSCDERIV_add NSCDERIV_CDERIV_iff [symmetric]) |
|
692 |
||
693 |
||
694 |
subsection{*Lemmas for Multiplication*} |
|
695 |
||
696 |
lemma lemma_nscderiv1: "((a::hcomplex)*b) - (c*d) = (b*(a - c)) + (c*(b - d))" |
|
697 |
by (simp add: right_diff_distrib) |
|
698 |
||
699 |
lemma lemma_nscderiv2: |
|
20563 | 700 |
"[| (x - y) / z = hcomplex_of_complex D + yb; z \<noteq> 0; |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
701 |
z : Infinitesimal; yb : Infinitesimal |] |
20563 | 702 |
==> x - y @= 0" |
14405 | 703 |
apply (frule_tac c1 = z in hcomplex_mult_right_cancel [THEN iffD2], assumption) |
20563 | 704 |
apply (erule_tac V = " (x - y) / z = hcomplex_of_complex D + yb" in thin_rl) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
705 |
apply (auto intro!: Infinitesimal_HFinite_mult2 HFinite_add |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
706 |
simp add: mem_infmal_iff [symmetric]) |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
707 |
apply (erule Infinitesimal_subset_HFinite [THEN subsetD]) |
14405 | 708 |
done |
709 |
||
710 |
lemma NSCDERIV_mult: |
|
711 |
"[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] |
|
712 |
==> NSCDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" |
|
20563 | 713 |
apply (auto simp add: NSCDERIV_NSCLIM_iff NSCLIM_def) |
14405 | 714 |
apply (auto dest!: spec |
20563 | 715 |
simp add: starfun_lambda_cancel lemma_nscderiv1) |
716 |
apply (simp (no_asm) add: add_divide_distrib diff_divide_distrib) |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
717 |
apply (drule bex_Infinitesimal_iff2 [THEN iffD2])+ |
20563 | 718 |
apply (auto simp add: times_divide_eq_right [symmetric] |
719 |
simp del: times_divide_eq) |
|
720 |
apply (drule_tac D = Db in lemma_nscderiv2, assumption+) |
|
721 |
apply (drule_tac |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
722 |
approx_minus_iff [THEN iffD2, THEN bex_Infinitesimal_iff2 [THEN iffD2]]) |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
723 |
apply (auto intro!: approx_add_mono1 simp add: left_distrib right_distrib mult_commute add_assoc) |
14405 | 724 |
apply (rule_tac b1 = "hcomplex_of_complex Db * hcomplex_of_complex (f x) " in add_commute [THEN subst]) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
725 |
apply (auto intro!: Infinitesimal_add_approx_self2 [THEN approx_sym] |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
726 |
Infinitesimal_add Infinitesimal_mult |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
727 |
Infinitesimal_hcomplex_of_complex_mult |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
728 |
Infinitesimal_hcomplex_of_complex_mult2 |
15013 | 729 |
simp add: add_assoc [symmetric]) |
14405 | 730 |
done |
731 |
||
732 |
lemma CDERIV_mult: |
|
733 |
"[| CDERIV f x :> Da; CDERIV g x :> Db |] |
|
734 |
==> CDERIV (%x. f x * g x) x :> (Da * g(x)) + (Db * f(x))" |
|
735 |
by (simp add: NSCDERIV_mult NSCDERIV_CDERIV_iff [symmetric]) |
|
736 |
||
737 |
lemma NSCDERIV_cmult: "NSCDERIV f x :> D ==> NSCDERIV (%x. c * f x) x :> c*D" |
|
738 |
apply (simp add: times_divide_eq_right [symmetric] NSCDERIV_NSCLIM_iff |
|
14469 | 739 |
minus_mult_right right_distrib [symmetric] diff_minus |
14405 | 740 |
del: times_divide_eq_right minus_mult_right [symmetric]) |
741 |
apply (erule NSCLIM_const [THEN NSCLIM_mult]) |
|
742 |
done |
|
743 |
||
744 |
lemma CDERIV_cmult: "CDERIV f x :> D ==> CDERIV (%x. c * f x) x :> c*D" |
|
745 |
by (simp add: NSCDERIV_cmult NSCDERIV_CDERIV_iff [symmetric]) |
|
746 |
||
747 |
lemma NSCDERIV_minus: "NSCDERIV f x :> D ==> NSCDERIV (%x. -(f x)) x :> -D" |
|
14469 | 748 |
apply (simp add: NSCDERIV_NSCLIM_iff diff_minus) |
14405 | 749 |
apply (rule_tac t = "f x" in minus_minus [THEN subst]) |
750 |
apply (simp (no_asm_simp) add: minus_add_distrib [symmetric] |
|
751 |
del: minus_add_distrib minus_minus) |
|
752 |
apply (erule NSCLIM_minus) |
|
753 |
done |
|
754 |
||
755 |
lemma CDERIV_minus: "CDERIV f x :> D ==> CDERIV (%x. -(f x)) x :> -D" |
|
756 |
by (simp add: NSCDERIV_minus NSCDERIV_CDERIV_iff [symmetric]) |
|
757 |
||
758 |
lemma NSCDERIV_add_minus: |
|
759 |
"[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] |
|
760 |
==> NSCDERIV (%x. f x + -g x) x :> Da + -Db" |
|
761 |
by (blast dest: NSCDERIV_add NSCDERIV_minus) |
|
762 |
||
763 |
lemma CDERIV_add_minus: |
|
764 |
"[| CDERIV f x :> Da; CDERIV g x :> Db |] |
|
765 |
==> CDERIV (%x. f x + -g x) x :> Da + -Db" |
|
766 |
by (blast dest: CDERIV_add CDERIV_minus) |
|
767 |
||
768 |
lemma NSCDERIV_diff: |
|
769 |
"[| NSCDERIV f x :> Da; NSCDERIV g x :> Db |] |
|
770 |
==> NSCDERIV (%x. f x - g x) x :> Da - Db" |
|
14469 | 771 |
by (simp add: diff_minus NSCDERIV_add_minus) |
14405 | 772 |
|
773 |
lemma CDERIV_diff: |
|
774 |
"[| CDERIV f x :> Da; CDERIV g x :> Db |] |
|
775 |
==> CDERIV (%x. f x - g x) x :> Da - Db" |
|
14469 | 776 |
by (simp add: diff_minus CDERIV_add_minus) |
14405 | 777 |
|
778 |
||
779 |
subsection{*Chain Rule*} |
|
780 |
||
781 |
(* lemmas *) |
|
782 |
lemma NSCDERIV_zero: |
|
783 |
"[| NSCDERIV g x :> D; |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
784 |
( *f* g) (hcomplex_of_complex(x) + xa) = hcomplex_of_complex(g x); |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
785 |
xa : Infinitesimal; xa \<noteq> 0 |
14405 | 786 |
|] ==> D = 0" |
787 |
apply (simp add: nscderiv_def) |
|
788 |
apply (drule bspec, auto) |
|
789 |
done |
|
790 |
||
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
791 |
lemma NSCDERIV_approx: |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
792 |
"[| NSCDERIV f x :> D; h: Infinitesimal; h \<noteq> 0 |] |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
793 |
==> ( *f* f) (hcomplex_of_complex(x) + h) - hcomplex_of_complex(f x) @= 0" |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
794 |
apply (simp add: nscderiv_def mem_infmal_iff [symmetric]) |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
795 |
apply (rule Infinitesimal_ratio) |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
796 |
apply (rule_tac [3] approx_hcomplex_of_complex_HFinite, auto) |
14405 | 797 |
done |
798 |
||
799 |
||
800 |
(*--------------------------------------------------*) |
|
801 |
(* from one version of differentiability *) |
|
802 |
(* *) |
|
803 |
(* f(x) - f(a) *) |
|
804 |
(* --------------- @= Db *) |
|
805 |
(* x - a *) |
|
806 |
(* -------------------------------------------------*) |
|
807 |
||
808 |
lemma NSCDERIVD1: |
|
809 |
"[| NSCDERIV f (g x) :> Da; |
|
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
810 |
( *f* g) (hcomplex_of_complex(x) + xa) \<noteq> hcomplex_of_complex (g x); |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
811 |
( *f* g) (hcomplex_of_complex(x) + xa) @= hcomplex_of_complex (g x)|] |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
812 |
==> (( *f* f) (( *f* g) (hcomplex_of_complex(x) + xa)) |
14405 | 813 |
- hcomplex_of_complex (f (g x))) / |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
814 |
(( *f* g) (hcomplex_of_complex(x) + xa) - hcomplex_of_complex (g x)) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
815 |
@= hcomplex_of_complex (Da)" |
14469 | 816 |
by (simp add: NSCDERIV_NSCLIM_iff2 NSCLIM_def) |
14405 | 817 |
|
818 |
(*--------------------------------------------------*) |
|
819 |
(* from other version of differentiability *) |
|
820 |
(* *) |
|
821 |
(* f(x + h) - f(x) *) |
|
822 |
(* ----------------- @= Db *) |
|
823 |
(* h *) |
|
824 |
(*--------------------------------------------------*) |
|
825 |
||
826 |
lemma NSCDERIVD2: |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
827 |
"[| NSCDERIV g x :> Db; xa: Infinitesimal; xa \<noteq> 0 |] |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
828 |
==> (( *f* g) (hcomplex_of_complex x + xa) - hcomplex_of_complex(g x)) / xa |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
829 |
@= hcomplex_of_complex (Db)" |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
830 |
by (simp add: NSCDERIV_NSCLIM_iff NSCLIM_def mem_infmal_iff starfun_lambda_cancel) |
14405 | 831 |
|
832 |
lemma lemma_complex_chain: "(z::hcomplex) \<noteq> 0 ==> x*y = (x*inverse(z))*(z*y)" |
|
833 |
by auto |
|
834 |
||
835 |
||
836 |
text{*Chain rule*} |
|
837 |
theorem NSCDERIV_chain: |
|
838 |
"[| NSCDERIV f (g x) :> Da; NSCDERIV g x :> Db |] |
|
839 |
==> NSCDERIV (f o g) x :> Da * Db" |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
840 |
apply (simp (no_asm_simp) add: NSCDERIV_NSCLIM_iff NSCLIM_def mem_infmal_iff [symmetric]) |
14405 | 841 |
apply safe |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
842 |
apply (frule_tac f = g in NSCDERIV_approx) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
843 |
apply (auto simp add: starfun_lambda_cancel2 starfun_o [symmetric]) |
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
844 |
apply (case_tac "( *f* g) (hcomplex_of_complex (x) + xa) = hcomplex_of_complex (g x) ") |
14405 | 845 |
apply (drule_tac g = g in NSCDERIV_zero) |
17300 | 846 |
apply (auto simp add: divide_inverse) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
847 |
apply (rule_tac z1 = "( *f* g) (hcomplex_of_complex (x) + xa) - hcomplex_of_complex (g x) " and y1 = "inverse xa" in lemma_complex_chain [THEN ssubst]) |
14405 | 848 |
apply (simp (no_asm_simp)) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
849 |
apply (rule approx_mult_hcomplex_of_complex) |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
850 |
apply (auto intro!: NSCDERIVD1 intro: approx_minus_iff [THEN iffD2] |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14405
diff
changeset
|
851 |
simp add: diff_minus [symmetric] divide_inverse [symmetric]) |
14405 | 852 |
apply (blast intro: NSCDERIVD2) |
853 |
done |
|
854 |
||
855 |
text{*standard version*} |
|
856 |
lemma CDERIV_chain: |
|
857 |
"[| CDERIV f (g x) :> Da; CDERIV g x :> Db |] |
|
858 |
==> CDERIV (f o g) x :> Da * Db" |
|
859 |
by (simp add: NSCDERIV_CDERIV_iff [symmetric] NSCDERIV_chain) |
|
860 |
||
861 |
lemma CDERIV_chain2: |
|
862 |
"[| CDERIV f (g x) :> Da; CDERIV g x :> Db |] |
|
863 |
==> CDERIV (%x. f (g x)) x :> Da * Db" |
|
864 |
by (auto dest: CDERIV_chain simp add: o_def) |
|
865 |
||
866 |
||
867 |
subsection{* Differentiation of Natural Number Powers*} |
|
868 |
||
869 |
lemma NSCDERIV_Id [simp]: "NSCDERIV (%x. x) x :> 1" |
|
15228 | 870 |
by (simp add: NSCDERIV_NSCLIM_iff NSCLIM_def divide_self del: divide_self_if) |
14405 | 871 |
|
872 |
lemma CDERIV_Id [simp]: "CDERIV (%x. x) x :> 1" |
|
873 |
by (simp add: NSCDERIV_CDERIV_iff [symmetric]) |
|
874 |
||
875 |
lemmas isContc_Id = CDERIV_Id [THEN CDERIV_isContc, standard] |
|
13957 | 876 |
|
14405 | 877 |
text{*derivative of linear multiplication*} |
878 |
lemma CDERIV_cmult_Id [simp]: "CDERIV (op * c) x :> c" |
|
879 |
by (cut_tac c = c and x = x in CDERIV_Id [THEN CDERIV_cmult], simp) |
|
880 |
||
881 |
lemma NSCDERIV_cmult_Id [simp]: "NSCDERIV (op * c) x :> c" |
|
882 |
by (simp add: NSCDERIV_CDERIV_iff) |
|
883 |
||
884 |
lemma CDERIV_pow [simp]: |
|
885 |
"CDERIV (%x. x ^ n) x :> (complex_of_real (real n)) * (x ^ (n - Suc 0))" |
|
886 |
apply (induct_tac "n") |
|
887 |
apply (drule_tac [2] CDERIV_Id [THEN CDERIV_mult]) |
|
15013 | 888 |
apply (auto simp add: left_distrib real_of_nat_Suc) |
14405 | 889 |
apply (case_tac "n") |
890 |
apply (auto simp add: mult_ac add_commute) |
|
891 |
done |
|
892 |
||
893 |
text{*Nonstandard version*} |
|
894 |
lemma NSCDERIV_pow: |
|
895 |
"NSCDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))" |
|
896 |
by (simp add: NSCDERIV_CDERIV_iff) |
|
897 |
||
898 |
lemma lemma_CDERIV_subst: |
|
899 |
"[|CDERIV f x :> D; D = E|] ==> CDERIV f x :> E" |
|
900 |
by auto |
|
901 |
||
902 |
(*used once, in NSCDERIV_inverse*) |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
903 |
lemma Infinitesimal_add_not_zero: |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
904 |
"[| h: Infinitesimal; x \<noteq> 0 |] ==> hcomplex_of_complex x + h \<noteq> 0" |
14405 | 905 |
apply clarify |
906 |
apply (drule equals_zero_I, auto) |
|
907 |
done |
|
908 |
||
909 |
text{*Can't relax the premise @{term "x \<noteq> 0"}: it isn't continuous at zero*} |
|
910 |
lemma NSCDERIV_inverse: |
|
911 |
"x \<noteq> 0 ==> NSCDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))" |
|
912 |
apply (simp add: nscderiv_def Ball_def, clarify) |
|
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
913 |
apply (frule Infinitesimal_add_not_zero [where x=x]) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
914 |
apply (auto simp add: starfun_inverse_inverse diff_minus |
14405 | 915 |
simp del: minus_mult_left [symmetric] minus_mult_right [symmetric]) |
17318
bc1c75855f3d
starfun, starset, and other functions on NS types are now polymorphic;
huffman
parents:
17300
diff
changeset
|
916 |
apply (simp add: add_commute numeral_2_eq_2 inverse_add |
14405 | 917 |
inverse_mult_distrib [symmetric] inverse_minus_eq [symmetric] |
918 |
add_ac mult_ac |
|
15013 | 919 |
del: inverse_minus_eq inverse_mult_distrib |
920 |
minus_mult_right [symmetric] minus_mult_left [symmetric]) |
|
14405 | 921 |
apply (simp only: mult_assoc [symmetric] right_distrib) |
20559
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
922 |
apply (rule_tac y = " inverse (- hcomplex_of_complex x * hcomplex_of_complex x) " in approx_trans) |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
923 |
apply (rule inverse_add_Infinitesimal_approx2) |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
924 |
apply (auto dest!: hcomplex_of_complex_HFinite_diff_Infinitesimal |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
925 |
intro: HFinite_mult |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
926 |
simp add: inverse_minus_eq [symmetric] HFinite_minus_iff) |
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
huffman
parents:
20552
diff
changeset
|
927 |
apply (rule Infinitesimal_HFinite_mult2, auto) |
14405 | 928 |
done |
929 |
||
930 |
lemma CDERIV_inverse: |
|
931 |
"x \<noteq> 0 ==> CDERIV (%x. inverse(x)) x :> (-(inverse x ^ 2))" |
|
932 |
by (simp add: NSCDERIV_inverse NSCDERIV_CDERIV_iff [symmetric] |
|
933 |
del: complexpow_Suc) |
|
934 |
||
935 |
||
936 |
subsection{*Derivative of Reciprocals (Function @{term inverse})*} |
|
937 |
||
938 |
lemma CDERIV_inverse_fun: |
|
939 |
"[| CDERIV f x :> d; f(x) \<noteq> 0 |] |
|
940 |
==> CDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))" |
|
941 |
apply (rule mult_commute [THEN subst]) |
|
15013 | 942 |
apply (simp add: minus_mult_left power_inverse |
943 |
del: complexpow_Suc minus_mult_left [symmetric]) |
|
14405 | 944 |
apply (fold o_def) |
945 |
apply (blast intro!: CDERIV_chain CDERIV_inverse) |
|
946 |
done |
|
947 |
||
948 |
lemma NSCDERIV_inverse_fun: |
|
949 |
"[| NSCDERIV f x :> d; f(x) \<noteq> 0 |] |
|
950 |
==> NSCDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))" |
|
951 |
by (simp add: NSCDERIV_CDERIV_iff CDERIV_inverse_fun del: complexpow_Suc) |
|
952 |
||
953 |
||
954 |
subsection{* Derivative of Quotient*} |
|
955 |
||
956 |
lemma CDERIV_quotient: |
|
957 |
"[| CDERIV f x :> d; CDERIV g x :> e; g(x) \<noteq> 0 |] |
|
958 |
==> CDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)" |
|
14469 | 959 |
apply (simp add: diff_minus) |
14405 | 960 |
apply (drule_tac f = g in CDERIV_inverse_fun) |
961 |
apply (drule_tac [2] CDERIV_mult, assumption+) |
|
19233
77ca20b0ed77
renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents:
17373
diff
changeset
|
962 |
apply (simp add: divide_inverse left_distrib power_inverse minus_mult_left |
14430
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14405
diff
changeset
|
963 |
mult_ac |
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14405
diff
changeset
|
964 |
del: minus_mult_right [symmetric] minus_mult_left [symmetric] |
5cb24165a2e1
new material from Avigad, and simplified treatment of division by 0
paulson
parents:
14405
diff
changeset
|
965 |
complexpow_Suc) |
14405 | 966 |
done |
967 |
||
968 |
lemma NSCDERIV_quotient: |
|
969 |
"[| NSCDERIV f x :> d; NSCDERIV g x :> e; g(x) \<noteq> 0 |] |
|
970 |
==> NSCDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)" |
|
971 |
by (simp add: NSCDERIV_CDERIV_iff CDERIV_quotient del: complexpow_Suc) |
|
972 |
||
973 |
||
974 |
subsection{*Caratheodory Formulation of Derivative at a Point: Standard Proof*} |
|
975 |
||
976 |
lemma CLIM_equal: |
|
977 |
"[| \<forall>x. x \<noteq> a --> (f x = g x) |] ==> (f -- a --C> l) = (g -- a --C> l)" |
|
978 |
by (simp add: CLIM_def complex_add_minus_iff) |
|
979 |
||
980 |
lemma CLIM_trans: |
|
981 |
"[| (%x. f(x) + -g(x)) -- a --C> 0; g -- a --C> l |] ==> f -- a --C> l" |
|
982 |
apply (drule CLIM_add, assumption) |
|
983 |
apply (simp add: complex_add_assoc) |
|
984 |
done |
|
985 |
||
986 |
lemma CARAT_CDERIV: |
|
987 |
"(CDERIV f x :> l) = |
|
988 |
(\<exists>g. (\<forall>z. f z - f x = g z * (z - x)) & isContc g x & g x = l)" |
|
989 |
apply safe |
|
990 |
apply (rule_tac x = "%z. if z=x then l else (f (z) - f (x)) / (z-x)" in exI) |
|
991 |
apply (auto simp add: mult_assoc isContc_iff CDERIV_iff) |
|
992 |
apply (rule_tac [!] CLIM_equal [THEN iffD1], auto) |
|
993 |
done |
|
994 |
||
995 |
||
996 |
lemma CARAT_NSCDERIV: |
|
997 |
"NSCDERIV f x :> l ==> |
|
998 |
\<exists>g. (\<forall>z. f z - f x = g z * (z - x)) & isNSContc g x & g x = l" |
|
14469 | 999 |
by (simp add: NSCDERIV_CDERIV_iff isNSContc_isContc_iff CARAT_CDERIV) |
14405 | 1000 |
|
1001 |
lemma CARAT_CDERIVD: |
|
1002 |
"(\<forall>z. f z - f x = g z * (z - x)) & isNSContc g x & g x = l |
|
1003 |
==> NSCDERIV f x :> l" |
|
20732 | 1004 |
by (auto simp add: NSCDERIV_iff2 isNSContc_def starfun_if_eq); |
15234
ec91a90c604e
simplification tweaks for better arithmetic reasoning
paulson
parents:
15228
diff
changeset
|
1005 |
|
14405 | 1006 |
end |