| author | huffman | 
| Tue, 08 May 2007 01:10:55 +0200 | |
| changeset 22854 | 51087b1cc77d | 
| parent 21839 | 54018ed3b99d | 
| child 22883 | 005be8dafce0 | 
| 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 | |
| 14405 | 36 | |
| 37 | subsection{*Limit of Complex to Complex Function*}
 | |
| 38 | ||
| 21792 | 39 | lemma NSLIM_Re: "f -- a --NS> L ==> (%x. Re(f x)) -- a --NS> Re(L)" | 
| 20659 
66b8455090b8
changed constants into abbreviations; shortened proofs
 huffman parents: 
20563diff
changeset | 40 | by (simp add: NSLIM_def starfunC_approx_Re_Im_iff | 
| 14469 | 41 | hRe_hcomplex_of_complex) | 
| 42 | ||
| 21792 | 43 | lemma NSLIM_Im: "f -- a --NS> L ==> (%x. Im(f x)) -- a --NS> Im(L)" | 
| 20659 
66b8455090b8
changed constants into abbreviations; shortened proofs
 huffman parents: 
20563diff
changeset | 44 | by (simp add: NSLIM_def starfunC_approx_Re_Im_iff | 
| 14469 | 45 | hIm_hcomplex_of_complex) | 
| 14405 | 46 | |
| 21792 | 47 | (** get this result easily now **) | 
| 48 | lemma LIM_Re: "f -- a --> L ==> (%x. Re(f x)) -- a --> Re(L)" | |
| 49 | by (simp add: LIM_NSLIM_iff NSLIM_Re) | |
| 14405 | 50 | |
| 21792 | 51 | lemma LIM_Im: "f -- a --> L ==> (%x. Im(f x)) -- a --> Im(L)" | 
| 52 | by (simp add: LIM_NSLIM_iff NSLIM_Im) | |
| 14405 | 53 | |
| 21792 | 54 | lemma LIM_cnj: "f -- a --> L ==> (%x. cnj (f x)) -- a --> cnj L" | 
| 55 | by (simp add: LIM_def complex_cnj_diff [symmetric]) | |
| 14405 | 56 | |
| 21792 | 57 | lemma LIM_cnj_iff: "((%x. cnj (f x)) -- a --> cnj L) = (f -- a --> L)" | 
| 58 | by (simp add: LIM_def complex_cnj_diff [symmetric]) | |
| 14405 | 59 | |
| 21792 | 60 | (*** NSLIM_not zero and hence LIM_not_zero ***) | 
| 14405 | 61 | |
| 21792 | 62 | lemma NSCLIM_not_zero: "k \<noteq> 0 ==> ~ ((%x::complex. k) -- x --NS> 0)" | 
| 63 | apply (auto simp del: star_of_zero simp add: NSLIM_def) | |
| 14405 | 64 | 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: 
20552diff
changeset | 65 | apply (auto intro: Infinitesimal_add_approx_self [THEN approx_sym] | 
| 17373 
27509e72f29e
removed duplicated lemmas; convert more proofs to transfer principle
 huffman parents: 
17318diff
changeset | 66 | simp del: star_of_zero) | 
| 14405 | 67 | done | 
| 68 | ||
| 21792 | 69 | (* [| k \<noteq> 0; (%x. k) -- x --NS> 0 |] ==> R *) | 
| 14405 | 70 | lemmas NSCLIM_not_zeroE = NSCLIM_not_zero [THEN notE, standard] | 
| 71 | ||
| 21792 | 72 | (*** NSLIM_const hence LIM_const ***) | 
| 14405 | 73 | |
| 21792 | 74 | lemma NSCLIM_const_eq: "(%x::complex. k) -- x --NS> L ==> k = L" | 
| 14405 | 75 | apply (rule ccontr) | 
| 21792 | 76 | apply (drule NSLIM_zero) | 
| 14405 | 77 | apply (rule NSCLIM_not_zeroE [of "k-L"], auto) | 
| 78 | done | |
| 79 | ||
| 21792 | 80 | (*** NSLIM and hence LIM are unique ***) | 
| 14405 | 81 | |
| 21792 | 82 | lemma NSCLIM_unique: "[| f -- (x::complex) --NS> L; f -- x --NS> M |] ==> L = M" | 
| 83 | apply (drule (1) NSLIM_diff) | |
| 84 | apply (drule NSLIM_minus) | |
| 14405 | 85 | apply (auto dest!: NSCLIM_const_eq [symmetric]) | 
| 86 | done | |
| 87 | ||
| 21831 | 88 | lemma starfun_norm: "( *f* (\<lambda>x. norm (f x))) = (\<lambda>x. hnorm (( *f* f) x))" | 
| 89 | by transfer (rule refl) | |
| 90 | ||
| 91 | lemma star_of_Re [simp]: "star_of (Re x) = hRe (star_of x)" | |
| 92 | by transfer (rule refl) | |
| 93 | ||
| 94 | lemma star_of_Im [simp]: "star_of (Im x) = hIm (star_of x)" | |
| 95 | by transfer (rule refl) | |
| 96 | ||
| 97 | text"" | |
| 14405 | 98 | (** another equivalence result **) | 
| 99 | lemma NSCLIM_NSCRLIM_iff: | |
| 21792 | 100 | "(f -- x --NS> L) = ((%y. cmod(f y - L)) -- x --NS> 0)" | 
| 21831 | 101 | by (simp add: NSLIM_def starfun_norm | 
| 102 | approx_approx_zero_iff [symmetric] approx_minus_iff [symmetric]) | |
| 14405 | 103 | |
| 104 | (** much, much easier standard proof **) | |
| 21792 | 105 | lemma CLIM_CRLIM_iff: "(f -- x --> L) = ((%y. cmod(f y - L)) -- x --> 0)" | 
| 106 | by (simp add: LIM_def) | |
| 14405 | 107 | |
| 108 | (* so this is nicer nonstandard proof *) | |
| 109 | lemma NSCLIM_NSCRLIM_iff2: | |
| 21792 | 110 | "(f -- x --NS> L) = ((%y. cmod(f y - L)) -- x --NS> 0)" | 
| 111 | by (simp add: LIM_NSLIM_iff [symmetric] CLIM_CRLIM_iff) | |
| 14405 | 112 | |
| 21792 | 113 | lemma NSLIM_NSCRLIM_Re_Im_iff: | 
| 114 | "(f -- a --NS> L) = ((%x. Re(f x)) -- a --NS> Re(L) & | |
| 115 | (%x. Im(f x)) -- a --NS> Im(L))" | |
| 116 | apply (auto intro: NSLIM_Re NSLIM_Im) | |
| 21831 | 117 | apply (auto simp add: NSLIM_def starfun_Re starfun_Im) | 
| 118 | apply (auto dest!: spec) | |
| 21839 
54018ed3b99d
added lemmas about hRe, hIm, HComplex; removed all uses of star_n
 huffman parents: 
21831diff
changeset | 119 | apply (simp add: hcomplex_approx_iff) | 
| 14405 | 120 | done | 
| 121 | ||
| 21792 | 122 | lemma LIM_CRLIM_Re_Im_iff: | 
| 123 | "(f -- a --> L) = ((%x. Re(f x)) -- a --> Re(L) & | |
| 124 | (%x. Im(f x)) -- a --> Im(L))" | |
| 125 | by (simp add: LIM_NSLIM_iff NSLIM_NSCRLIM_Re_Im_iff) | |
| 14405 | 126 | |
| 127 | ||
| 128 | subsection{*Continuity*}
 | |
| 129 | ||
| 21792 | 130 | lemma NSLIM_isContc_iff: | 
| 131 | "(f -- a --NS> f a) = ((%h. f(a + h)) -- 0 --NS> f a)" | |
| 20659 
66b8455090b8
changed constants into abbreviations; shortened proofs
 huffman parents: 
20563diff
changeset | 132 | by (rule NSLIM_h_iff) | 
| 14405 | 133 | |
| 134 | subsection{*Functions from Complex to Reals*}
 | |
| 135 | ||
| 21792 | 136 | lemma isNSContCR_cmod [simp]: "isNSCont cmod (a)" | 
| 20559 
2116b7a371c7
removed capprox, CFinite, CInfinite, CInfinitesimal, cmonad, and cgalaxy in favor of polymorphic constants
 huffman parents: 
20552diff
changeset | 137 | by (auto intro: approx_hcmod_approx | 
| 14405 | 138 | simp add: starfunCR_cmod hcmod_hcomplex_of_complex [symmetric] | 
| 21792 | 139 | isNSCont_def) | 
| 14405 | 140 | |
| 21792 | 141 | lemma isContCR_cmod [simp]: "isCont cmod (a)" | 
| 142 | by (simp add: isNSCont_isCont_iff [symmetric]) | |
| 14405 | 143 | |
| 21792 | 144 | lemma isCont_Re: "isCont f a ==> isCont (%x. Re (f x)) a" | 
| 145 | by (simp add: isCont_def LIM_Re) | |
| 14405 | 146 | |
| 21792 | 147 | lemma isCont_Im: "isCont f a ==> isCont (%x. Im (f x)) a" | 
| 148 | by (simp add: isCont_def LIM_Im) | |
| 14405 | 149 | |
| 150 | subsection{* Differentiation of Natural Number Powers*}
 | |
| 151 | ||
| 152 | lemma CDERIV_pow [simp]: | |
| 21792 | 153 | "DERIV (%x. x ^ n) x :> (complex_of_real (real n)) * (x ^ (n - Suc 0))" | 
| 14405 | 154 | apply (induct_tac "n") | 
| 21792 | 155 | apply (drule_tac [2] DERIV_Id [THEN DERIV_mult]) | 
| 15013 | 156 | apply (auto simp add: left_distrib real_of_nat_Suc) | 
| 14405 | 157 | apply (case_tac "n") | 
| 158 | apply (auto simp add: mult_ac add_commute) | |
| 159 | done | |
| 160 | ||
| 161 | text{*Nonstandard version*}
 | |
| 162 | lemma NSCDERIV_pow: | |
| 21792 | 163 | "NSDERIV (%x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))" | 
| 164 | by (simp add: NSDERIV_DERIV_iff) | |
| 14405 | 165 | |
| 166 | text{*Can't relax the premise @{term "x \<noteq> 0"}: it isn't continuous at zero*}
 | |
| 167 | lemma NSCDERIV_inverse: | |
| 21792 | 168 | "(x::complex) \<noteq> 0 ==> NSDERIV (%x. inverse(x)) x :> (- (inverse x ^ 2))" | 
| 169 | unfolding numeral_2_eq_2 | |
| 170 | by (rule NSDERIV_inverse) | |
| 14405 | 171 | |
| 172 | lemma CDERIV_inverse: | |
| 21792 | 173 | "(x::complex) \<noteq> 0 ==> DERIV (%x. inverse(x)) x :> (-(inverse x ^ 2))" | 
| 174 | unfolding numeral_2_eq_2 | |
| 175 | by (rule DERIV_inverse) | |
| 14405 | 176 | |
| 177 | ||
| 178 | subsection{*Derivative of Reciprocals (Function @{term inverse})*}
 | |
| 179 | ||
| 180 | lemma CDERIV_inverse_fun: | |
| 21792 | 181 | "[| DERIV f x :> d; f(x) \<noteq> (0::complex) |] | 
| 182 | ==> DERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))" | |
| 183 | unfolding numeral_2_eq_2 | |
| 184 | by (rule DERIV_inverse_fun) | |
| 14405 | 185 | |
| 186 | lemma NSCDERIV_inverse_fun: | |
| 21792 | 187 | "[| NSDERIV f x :> d; f(x) \<noteq> (0::complex) |] | 
| 188 | ==> NSDERIV (%x. inverse(f x)) x :> (- (d * inverse(f(x) ^ 2)))" | |
| 189 | unfolding numeral_2_eq_2 | |
| 190 | by (rule NSDERIV_inverse_fun) | |
| 14405 | 191 | |
| 192 | ||
| 193 | subsection{* Derivative of Quotient*}
 | |
| 194 | ||
| 195 | lemma CDERIV_quotient: | |
| 21792 | 196 | "[| DERIV f x :> d; DERIV g x :> e; g(x) \<noteq> (0::complex) |] | 
| 197 | ==> DERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)" | |
| 198 | unfolding numeral_2_eq_2 | |
| 199 | by (rule DERIV_quotient) | |
| 14405 | 200 | |
| 201 | lemma NSCDERIV_quotient: | |
| 21792 | 202 | "[| NSDERIV f x :> d; NSDERIV g x :> e; g(x) \<noteq> (0::complex) |] | 
| 203 | ==> NSDERIV (%y. f(y) / (g y)) x :> (d*g(x) - (e*f(x))) / (g(x) ^ 2)" | |
| 204 | unfolding numeral_2_eq_2 | |
| 205 | by (rule NSDERIV_quotient) | |
| 14405 | 206 | |
| 207 | ||
| 208 | subsection{*Caratheodory Formulation of Derivative at a Point: Standard Proof*}
 | |
| 209 | ||
| 210 | lemma CARAT_CDERIVD: | |
| 21792 | 211 | "(\<forall>z. f z - f x = g z * (z - x)) & isNSCont g x & g x = l | 
| 212 | ==> NSDERIV f x :> l" | |
| 213 | by clarify (rule CARAT_DERIVD) | |
| 15234 
ec91a90c604e
simplification tweaks for better arithmetic reasoning
 paulson parents: 
15228diff
changeset | 214 | |
| 14405 | 215 | end |