author | wenzelm |
Fri, 29 Nov 2024 17:40:15 +0100 | |
changeset 81507 | 08574da77b4a |
parent 75866 | 9eeed5c424f9 |
permissions | -rw-r--r-- |
62479 | 1 |
(* Title: HOL/Nonstandard_Analysis/CLim.thy |
2 |
Author: Jacques D. Fleuriot |
|
3 |
Copyright: 2001 University of Edinburgh |
|
27468 | 4 |
Conversion to Isar and new proofs by Lawrence C Paulson, 2004 |
5 |
*) |
|
6 |
||
64601 | 7 |
section \<open>Limits, Continuity and Differentiation for Complex Functions\<close> |
27468 | 8 |
|
9 |
theory CLim |
|
64601 | 10 |
imports CStar |
27468 | 11 |
begin |
12 |
||
13 |
(*not in simpset?*) |
|
70723
4e39d87c9737
imported new material mostly due to Sébastien Gouëzel
paulson <lp15@cam.ac.uk>
parents:
70228
diff
changeset
|
14 |
declare epsilon_not_zero [simp] |
27468 | 15 |
|
16 |
(*??generalize*) |
|
64601 | 17 |
lemma lemma_complex_mult_inverse_squared [simp]: "x \<noteq> 0 \<Longrightarrow> x * (inverse x)\<^sup>2 = inverse x" |
18 |
for x :: complex |
|
19 |
by (simp add: numeral_2_eq_2) |
|
20 |
||
21 |
text \<open>Changing the quantified variable. Install earlier?\<close> |
|
22 |
lemma all_shift: "(\<forall>x::'a::comm_ring_1. P x) \<longleftrightarrow> (\<forall>x. P (x - a))" |
|
75866 | 23 |
by (metis add_diff_cancel) |
27468 | 24 |
|
64601 | 25 |
subsection \<open>Limit of Complex to Complex Function\<close> |
26 |
||
27 |
lemma NSLIM_Re: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Re L" |
|
28 |
by (simp add: NSLIM_def starfunC_approx_Re_Im_iff hRe_hcomplex_of_complex) |
|
27468 | 29 |
|
64601 | 30 |
lemma NSLIM_Im: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<Longrightarrow> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Im L" |
31 |
by (simp add: NSLIM_def starfunC_approx_Re_Im_iff hIm_hcomplex_of_complex) |
|
27468 | 32 |
|
64601 | 33 |
lemma LIM_Re: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow> Re L" |
34 |
for f :: "'a::real_normed_vector \<Rightarrow> complex" |
|
35 |
by (simp add: LIM_NSLIM_iff NSLIM_Re) |
|
27468 | 36 |
|
64601 | 37 |
lemma LIM_Im: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow> Im L" |
38 |
for f :: "'a::real_normed_vector \<Rightarrow> complex" |
|
39 |
by (simp add: LIM_NSLIM_iff NSLIM_Im) |
|
27468 | 40 |
|
64601 | 41 |
lemma LIM_cnj: "f \<midarrow>a\<rightarrow> L \<Longrightarrow> (\<lambda>x. cnj (f x)) \<midarrow>a\<rightarrow> cnj L" |
42 |
for f :: "'a::real_normed_vector \<Rightarrow> complex" |
|
43 |
by (simp add: LIM_eq complex_cnj_diff [symmetric] del: complex_cnj_diff) |
|
27468 | 44 |
|
64601 | 45 |
lemma LIM_cnj_iff: "((\<lambda>x. cnj (f x)) \<midarrow>a\<rightarrow> cnj L) \<longleftrightarrow> f \<midarrow>a\<rightarrow> L" |
46 |
for f :: "'a::real_normed_vector \<Rightarrow> complex" |
|
47 |
by (simp add: LIM_eq complex_cnj_diff [symmetric] del: complex_cnj_diff) |
|
27468 | 48 |
|
49 |
lemma starfun_norm: "( *f* (\<lambda>x. norm (f x))) = (\<lambda>x. hnorm (( *f* f) x))" |
|
64601 | 50 |
by transfer (rule refl) |
27468 | 51 |
|
52 |
lemma star_of_Re [simp]: "star_of (Re x) = hRe (star_of x)" |
|
64601 | 53 |
by transfer (rule refl) |
27468 | 54 |
|
55 |
lemma star_of_Im [simp]: "star_of (Im x) = hIm (star_of x)" |
|
64601 | 56 |
by transfer (rule refl) |
27468 | 57 |
|
64601 | 58 |
text \<open>Another equivalence result.\<close> |
59 |
lemma NSCLIM_NSCRLIM_iff: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>y. cmod (f y - L)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0" |
|
60 |
by (simp add: NSLIM_def starfun_norm |
|
61 |
approx_approx_zero_iff [symmetric] approx_minus_iff [symmetric]) |
|
27468 | 62 |
|
64601 | 63 |
text \<open>Much, much easier standard proof.\<close> |
64 |
lemma CLIM_CRLIM_iff: "f \<midarrow>x\<rightarrow> L \<longleftrightarrow> (\<lambda>y. cmod (f y - L)) \<midarrow>x\<rightarrow> 0" |
|
65 |
for f :: "'a::real_normed_vector \<Rightarrow> complex" |
|
66 |
by (simp add: LIM_eq) |
|
27468 | 67 |
|
64601 | 68 |
text \<open>So this is nicer nonstandard proof.\<close> |
69 |
lemma NSCLIM_NSCRLIM_iff2: "f \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>y. cmod (f y - L)) \<midarrow>x\<rightarrow>\<^sub>N\<^sub>S 0" |
|
70 |
by (simp add: LIM_NSLIM_iff [symmetric] CLIM_CRLIM_iff) |
|
27468 | 71 |
|
72 |
lemma NSLIM_NSCRLIM_Re_Im_iff: |
|
64601 | 73 |
"f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S L \<longleftrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Re L \<and> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S Im L" |
74 |
apply (auto intro: NSLIM_Re NSLIM_Im) |
|
75 |
apply (auto simp add: NSLIM_def starfun_Re starfun_Im) |
|
76 |
apply (auto dest!: spec) |
|
77 |
apply (simp add: hcomplex_approx_iff) |
|
78 |
done |
|
79 |
||
80 |
lemma LIM_CRLIM_Re_Im_iff: "f \<midarrow>a\<rightarrow> L \<longleftrightarrow> (\<lambda>x. Re (f x)) \<midarrow>a\<rightarrow> Re L \<and> (\<lambda>x. Im (f x)) \<midarrow>a\<rightarrow> Im L" |
|
81 |
for f :: "'a::real_normed_vector \<Rightarrow> complex" |
|
82 |
by (simp add: LIM_NSLIM_iff NSLIM_NSCRLIM_Re_Im_iff) |
|
83 |
||
84 |
||
85 |
subsection \<open>Continuity\<close> |
|
86 |
||
87 |
lemma NSLIM_isContc_iff: "f \<midarrow>a\<rightarrow>\<^sub>N\<^sub>S f a \<longleftrightarrow> (\<lambda>h. f (a + h)) \<midarrow>0\<rightarrow>\<^sub>N\<^sub>S f a" |
|
70228
2d5b122aa0ff
De-applying and combining lemmas to make structured proofs
paulson <lp15@cam.ac.uk>
parents:
69597
diff
changeset
|
88 |
by (rule NSLIM_at0_iff) |
27468 | 89 |
|
64601 | 90 |
|
91 |
subsection \<open>Functions from Complex to Reals\<close> |
|
92 |
||
93 |
lemma isNSContCR_cmod [simp]: "isNSCont cmod a" |
|
94 |
by (auto intro: approx_hnorm |
|
95 |
simp: starfunCR_cmod hcmod_hcomplex_of_complex [symmetric] isNSCont_def) |
|
96 |
||
97 |
lemma isContCR_cmod [simp]: "isCont cmod a" |
|
98 |
by (simp add: isNSCont_isCont_iff [symmetric]) |
|
99 |
||
100 |
lemma isCont_Re: "isCont f a \<Longrightarrow> isCont (\<lambda>x. Re (f x)) a" |
|
101 |
for f :: "'a::real_normed_vector \<Rightarrow> complex" |
|
102 |
by (simp add: isCont_def LIM_Re) |
|
103 |
||
104 |
lemma isCont_Im: "isCont f a \<Longrightarrow> isCont (\<lambda>x. Im (f x)) a" |
|
105 |
for f :: "'a::real_normed_vector \<Rightarrow> complex" |
|
106 |
by (simp add: isCont_def LIM_Im) |
|
27468 | 107 |
|
108 |
||
64601 | 109 |
subsection \<open>Differentiation of Natural Number Powers\<close> |
27468 | 110 |
|
64601 | 111 |
lemma CDERIV_pow [simp]: "DERIV (\<lambda>x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - Suc 0))" |
112 |
apply (induct n) |
|
113 |
apply (drule_tac [2] DERIV_ident [THEN DERIV_mult]) |
|
114 |
apply (auto simp add: distrib_right of_nat_Suc) |
|
115 |
apply (case_tac "n") |
|
116 |
apply (auto simp add: ac_simps) |
|
117 |
done |
|
27468 | 118 |
|
64601 | 119 |
text \<open>Nonstandard version.\<close> |
120 |
lemma NSCDERIV_pow: "NSDERIV (\<lambda>x. x ^ n) x :> complex_of_real (real n) * (x ^ (n - 1))" |
|
121 |
by (metis CDERIV_pow NSDERIV_DERIV_iff One_nat_def) |
|
27468 | 122 |
|
69597 | 123 |
text \<open>Can't relax the premise \<^term>\<open>x \<noteq> 0\<close>: it isn't continuous at zero.\<close> |
64601 | 124 |
lemma NSCDERIV_inverse: "x \<noteq> 0 \<Longrightarrow> NSDERIV (\<lambda>x. inverse x) x :> - (inverse x)\<^sup>2" |
125 |
for x :: complex |
|
126 |
unfolding numeral_2_eq_2 by (rule NSDERIV_inverse) |
|
27468 | 127 |
|
64601 | 128 |
lemma CDERIV_inverse: "x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. inverse x) x :> - (inverse x)\<^sup>2" |
129 |
for x :: complex |
|
130 |
unfolding numeral_2_eq_2 by (rule DERIV_inverse) |
|
27468 | 131 |
|
132 |
||
69597 | 133 |
subsection \<open>Derivative of Reciprocals (Function \<^term>\<open>inverse\<close>)\<close> |
27468 | 134 |
|
135 |
lemma CDERIV_inverse_fun: |
|
64601 | 136 |
"DERIV f x :> d \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> DERIV (\<lambda>x. inverse (f x)) x :> - (d * inverse ((f x)\<^sup>2))" |
137 |
for x :: complex |
|
138 |
unfolding numeral_2_eq_2 by (rule DERIV_inverse_fun) |
|
27468 | 139 |
|
140 |
lemma NSCDERIV_inverse_fun: |
|
64601 | 141 |
"NSDERIV f x :> d \<Longrightarrow> f x \<noteq> 0 \<Longrightarrow> NSDERIV (\<lambda>x. inverse (f x)) x :> - (d * inverse ((f x)\<^sup>2))" |
142 |
for x :: complex |
|
143 |
unfolding numeral_2_eq_2 by (rule NSDERIV_inverse_fun) |
|
27468 | 144 |
|
145 |
||
64601 | 146 |
subsection \<open>Derivative of Quotient\<close> |
27468 | 147 |
|
148 |
lemma CDERIV_quotient: |
|
64601 | 149 |
"DERIV f x :> d \<Longrightarrow> DERIV g x :> e \<Longrightarrow> g(x) \<noteq> 0 \<Longrightarrow> |
150 |
DERIV (\<lambda>y. f y / g y) x :> (d * g x - (e * f x)) / (g x)\<^sup>2" |
|
151 |
for x :: complex |
|
152 |
unfolding numeral_2_eq_2 by (rule DERIV_quotient) |
|
27468 | 153 |
|
154 |
lemma NSCDERIV_quotient: |
|
64601 | 155 |
"NSDERIV f x :> d \<Longrightarrow> NSDERIV g x :> e \<Longrightarrow> g x \<noteq> (0::complex) \<Longrightarrow> |
156 |
NSDERIV (\<lambda>y. f y / g y) x :> (d * g x - (e * f x)) / (g x)\<^sup>2" |
|
157 |
unfolding numeral_2_eq_2 by (rule NSDERIV_quotient) |
|
27468 | 158 |
|
159 |
||
64601 | 160 |
subsection \<open>Caratheodory Formulation of Derivative at a Point: Standard Proof\<close> |
27468 | 161 |
|
162 |
lemma CARAT_CDERIVD: |
|
64601 | 163 |
"(\<forall>z. f z - f x = g z * (z - x)) \<and> isNSCont g x \<and> g x = l \<Longrightarrow> NSDERIV f x :> l" |
164 |
by clarify (rule CARAT_DERIVD) |
|
27468 | 165 |
|
166 |
end |