| author | blanchet | 
| Tue, 03 Mar 2015 16:37:45 +0100 | |
| changeset 59575 | 55f5e1cbf2a7 | 
| parent 59554 | 4044f53326c9 | 
| child 59615 | fdfdf89a83a6 | 
| permissions | -rw-r--r-- | 
| 56215 | 1  | 
(* Author: John Harrison, Marco Maggesi, Graziano Gentili, Gianni Ciolli, Valentina Bruno  | 
2  | 
Ported from "hol_light/Multivariate/canal.ml" by L C Paulson (2014)  | 
|
3  | 
*)  | 
|
4  | 
||
| 58877 | 5  | 
section {* Complex Analysis Basics *}
 | 
| 56215 | 6  | 
|
7  | 
theory Complex_Analysis_Basics  | 
|
8  | 
imports "~~/src/HOL/Multivariate_Analysis/Cartesian_Euclidean_Space"  | 
|
9  | 
begin  | 
|
10  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
11  | 
subsection{*General lemmas*}
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
12  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
13  | 
lemma has_derivative_mult_right:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
14  | 
fixes c:: "'a :: real_normed_algebra"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
15  | 
shows "((op * c) has_derivative (op * c)) F"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
16  | 
by (rule has_derivative_mult_right [OF has_derivative_id])  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
17  | 
|
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
18  | 
lemma has_derivative_of_real[derivative_intros, simp]:  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
19  | 
"(f has_derivative f') F \<Longrightarrow> ((\<lambda>x. of_real (f x)) has_derivative (\<lambda>x. of_real (f' x))) F"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
20  | 
using bounded_linear.has_derivative[OF bounded_linear_of_real] .  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
21  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
22  | 
lemma has_vector_derivative_real_complex:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
23  | 
"DERIV f (of_real a) :> f' \<Longrightarrow> ((\<lambda>x. f (of_real x)) has_vector_derivative f') (at a)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
24  | 
using has_derivative_compose[of of_real of_real a UNIV f "op * f'"]  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
25  | 
by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def)  | 
| 56215 | 26  | 
|
| 
56238
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
27  | 
lemma fact_cancel:  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
28  | 
fixes c :: "'a::real_field"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
29  | 
shows "of_nat (Suc n) * c / of_nat (fact (Suc n)) = c / of_nat (fact n)"  | 
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
30  | 
by (simp add: of_nat_mult del: of_nat_Suc times_nat.simps)  | 
| 
56889
 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 
hoelzl 
parents: 
56479 
diff
changeset
 | 
31  | 
|
| 56215 | 32  | 
lemma linear_times:  | 
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
33  | 
fixes c::"'a::real_algebra" shows "linear (\<lambda>x. c * x)"  | 
| 56215 | 34  | 
by (auto simp: linearI distrib_left)  | 
35  | 
||
36  | 
lemma bilinear_times:  | 
|
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
37  | 
fixes c::"'a::real_algebra" shows "bilinear (\<lambda>x y::'a. x*y)"  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
38  | 
by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI)  | 
| 56215 | 39  | 
|
40  | 
lemma linear_cnj: "linear cnj"  | 
|
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
41  | 
using bounded_linear.linear[OF bounded_linear_cnj] .  | 
| 56215 | 42  | 
|
43  | 
lemma tendsto_mult_left:  | 
|
44  | 
fixes c::"'a::real_normed_algebra"  | 
|
45  | 
shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. c * (f x)) ---> c * l) F"  | 
|
46  | 
by (rule tendsto_mult [OF tendsto_const])  | 
|
47  | 
||
48  | 
lemma tendsto_mult_right:  | 
|
49  | 
fixes c::"'a::real_normed_algebra"  | 
|
50  | 
shows "(f ---> l) F \<Longrightarrow> ((\<lambda>x. (f x) * c) ---> l * c) F"  | 
|
51  | 
by (rule tendsto_mult [OF _ tendsto_const])  | 
|
52  | 
||
53  | 
lemma tendsto_Re_upper:  | 
|
54  | 
assumes "~ (trivial_limit F)"  | 
|
55  | 
"(f ---> l) F"  | 
|
56  | 
"eventually (\<lambda>x. Re(f x) \<le> b) F"  | 
|
57  | 
shows "Re(l) \<le> b"  | 
|
58  | 
by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Re)  | 
|
59  | 
||
60  | 
lemma tendsto_Re_lower:  | 
|
61  | 
assumes "~ (trivial_limit F)"  | 
|
62  | 
"(f ---> l) F"  | 
|
63  | 
"eventually (\<lambda>x. b \<le> Re(f x)) F"  | 
|
64  | 
shows "b \<le> Re(l)"  | 
|
65  | 
by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Re)  | 
|
66  | 
||
67  | 
lemma tendsto_Im_upper:  | 
|
68  | 
assumes "~ (trivial_limit F)"  | 
|
69  | 
"(f ---> l) F"  | 
|
70  | 
"eventually (\<lambda>x. Im(f x) \<le> b) F"  | 
|
71  | 
shows "Im(l) \<le> b"  | 
|
72  | 
by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Im)  | 
|
73  | 
||
74  | 
lemma tendsto_Im_lower:  | 
|
75  | 
assumes "~ (trivial_limit F)"  | 
|
76  | 
"(f ---> l) F"  | 
|
77  | 
"eventually (\<lambda>x. b \<le> Im(f x)) F"  | 
|
78  | 
shows "b \<le> Im(l)"  | 
|
79  | 
by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Im)  | 
|
80  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
81  | 
lemma lambda_zero: "(\<lambda>h::'a::mult_zero. 0) = op * 0"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
82  | 
by auto  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
83  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
84  | 
lemma lambda_one: "(\<lambda>x::'a::monoid_mult. x) = op * 1"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
85  | 
by auto  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
86  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
87  | 
lemma has_real_derivative:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
88  | 
fixes f :: "real \<Rightarrow> real"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
89  | 
assumes "(f has_derivative f') F"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
90  | 
obtains c where "(f has_real_derivative c) F"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
91  | 
proof -  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
92  | 
obtain c where "f' = (\<lambda>x. x * c)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
93  | 
by (metis assms has_derivative_bounded_linear real_bounded_linear)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
94  | 
then show ?thesis  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
95  | 
by (metis assms that has_field_derivative_def mult_commute_abs)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
96  | 
qed  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
97  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
98  | 
lemma has_real_derivative_iff:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
99  | 
fixes f :: "real \<Rightarrow> real"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
100  | 
shows "(\<exists>c. (f has_real_derivative c) F) = (\<exists>D. (f has_derivative D) F)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
101  | 
by (metis has_field_derivative_def has_real_derivative)  | 
| 56215 | 102  | 
|
103  | 
lemma continuous_mult_left:  | 
|
104  | 
fixes c::"'a::real_normed_algebra"  | 
|
105  | 
shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. c * f x)"  | 
|
106  | 
by (rule continuous_mult [OF continuous_const])  | 
|
107  | 
||
108  | 
lemma continuous_mult_right:  | 
|
109  | 
fixes c::"'a::real_normed_algebra"  | 
|
110  | 
shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x * c)"  | 
|
111  | 
by (rule continuous_mult [OF _ continuous_const])  | 
|
112  | 
||
113  | 
lemma continuous_on_mult_left:  | 
|
114  | 
fixes c::"'a::real_normed_algebra"  | 
|
115  | 
shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. c * f x)"  | 
|
116  | 
by (rule continuous_on_mult [OF continuous_on_const])  | 
|
117  | 
||
118  | 
lemma continuous_on_mult_right:  | 
|
119  | 
fixes c::"'a::real_normed_algebra"  | 
|
120  | 
shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x * c)"  | 
|
121  | 
by (rule continuous_on_mult [OF _ continuous_on_const])  | 
|
122  | 
||
| 
56371
 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 
hoelzl 
parents: 
56370 
diff
changeset
 | 
123  | 
lemma uniformly_continuous_on_cmul_right [continuous_intros]:  | 
| 56215 | 124  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"  | 
| 56332 | 125  | 
shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. f x * c)"  | 
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
126  | 
using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] .  | 
| 56215 | 127  | 
|
| 
56371
 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 
hoelzl 
parents: 
56370 
diff
changeset
 | 
128  | 
lemma uniformly_continuous_on_cmul_left[continuous_intros]:  | 
| 56215 | 129  | 
fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra"  | 
130  | 
assumes "uniformly_continuous_on s f"  | 
|
131  | 
shows "uniformly_continuous_on s (\<lambda>x. c * f x)"  | 
|
132  | 
by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right)  | 
|
133  | 
||
134  | 
lemma continuous_within_norm_id [continuous_intros]: "continuous (at x within S) norm"  | 
|
135  | 
by (rule continuous_norm [OF continuous_ident])  | 
|
136  | 
||
137  | 
lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm"  | 
|
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
138  | 
by (intro continuous_on_id continuous_on_norm)  | 
| 56215 | 139  | 
|
140  | 
subsection{*DERIV stuff*}
 | 
|
141  | 
||
142  | 
lemma DERIV_zero_connected_constant:  | 
|
143  | 
  fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
 | 
|
144  | 
assumes "connected s"  | 
|
145  | 
and "open s"  | 
|
146  | 
and "finite k"  | 
|
147  | 
and "continuous_on s f"  | 
|
148  | 
and "\<forall>x\<in>(s - k). DERIV f x :> 0"  | 
|
149  | 
obtains c where "\<And>x. x \<in> s \<Longrightarrow> f(x) = c"  | 
|
150  | 
using has_derivative_zero_connected_constant [OF assms(1-4)] assms  | 
|
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
151  | 
by (metis DERIV_const has_derivative_const Diff_iff at_within_open frechet_derivative_at has_field_derivative_def)  | 
| 56215 | 152  | 
|
153  | 
lemma DERIV_zero_constant:  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
154  | 
  fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a"
 | 
| 56215 | 155  | 
shows "\<lbrakk>convex s;  | 
156  | 
\<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)\<rbrakk>  | 
|
157  | 
\<Longrightarrow> \<exists>c. \<forall>x \<in> s. f(x) = c"  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
158  | 
by (auto simp: has_field_derivative_def lambda_zero intro: has_derivative_zero_constant)  | 
| 56215 | 159  | 
|
160  | 
lemma DERIV_zero_unique:  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
161  | 
  fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a"
 | 
| 56215 | 162  | 
assumes "convex s"  | 
163  | 
and d0: "\<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)"  | 
|
164  | 
and "a \<in> s"  | 
|
165  | 
and "x \<in> s"  | 
|
166  | 
shows "f x = f a"  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
167  | 
by (rule has_derivative_zero_unique [OF assms(1) _ assms(4,3)])  | 
| 56332 | 168  | 
(metis d0 has_field_derivative_imp_has_derivative lambda_zero)  | 
| 56215 | 169  | 
|
170  | 
lemma DERIV_zero_connected_unique:  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
171  | 
  fixes f :: "'a::{real_normed_field, real_inner} \<Rightarrow> 'a"
 | 
| 56215 | 172  | 
assumes "connected s"  | 
173  | 
and "open s"  | 
|
174  | 
and d0: "\<And>x. x\<in>s \<Longrightarrow> DERIV f x :> 0"  | 
|
175  | 
and "a \<in> s"  | 
|
176  | 
and "x \<in> s"  | 
|
177  | 
shows "f x = f a"  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
178  | 
by (rule has_derivative_zero_unique_connected [OF assms(2,1) _ assms(5,4)])  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
179  | 
(metis has_field_derivative_def lambda_zero d0)  | 
| 56215 | 180  | 
|
181  | 
lemma DERIV_transform_within:  | 
|
182  | 
assumes "(f has_field_derivative f') (at a within s)"  | 
|
183  | 
and "0 < d" "a \<in> s"  | 
|
184  | 
and "\<And>x. x\<in>s \<Longrightarrow> dist x a < d \<Longrightarrow> f x = g x"  | 
|
185  | 
shows "(g has_field_derivative f') (at a within s)"  | 
|
186  | 
using assms unfolding has_field_derivative_def  | 
|
| 56332 | 187  | 
by (blast intro: has_derivative_transform_within)  | 
| 56215 | 188  | 
|
189  | 
lemma DERIV_transform_within_open:  | 
|
190  | 
assumes "DERIV f a :> f'"  | 
|
191  | 
and "open s" "a \<in> s"  | 
|
192  | 
and "\<And>x. x\<in>s \<Longrightarrow> f x = g x"  | 
|
193  | 
shows "DERIV g a :> f'"  | 
|
194  | 
using assms unfolding has_field_derivative_def  | 
|
195  | 
by (metis has_derivative_transform_within_open)  | 
|
196  | 
||
197  | 
lemma DERIV_transform_at:  | 
|
198  | 
assumes "DERIV f a :> f'"  | 
|
199  | 
and "0 < d"  | 
|
200  | 
and "\<And>x. dist x a < d \<Longrightarrow> f x = g x"  | 
|
201  | 
shows "DERIV g a :> f'"  | 
|
202  | 
by (blast intro: assms DERIV_transform_within)  | 
|
203  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
204  | 
subsection {*Some limit theorems about real part of real series etc.*}
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
205  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
206  | 
(*MOVE? But not to Finite_Cartesian_Product*)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
207  | 
lemma sums_vec_nth :  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
208  | 
assumes "f sums a"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
209  | 
shows "(\<lambda>x. f x $ i) sums a $ i"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
210  | 
using assms unfolding sums_def  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
211  | 
by (auto dest: tendsto_vec_nth [where i=i])  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
212  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
213  | 
lemma summable_vec_nth :  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
214  | 
assumes "summable f"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
215  | 
shows "summable (\<lambda>x. f x $ i)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
216  | 
using assms unfolding summable_def  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
217  | 
by (blast intro: sums_vec_nth)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
218  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
219  | 
subsection {*Complex number lemmas *}
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
220  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
221  | 
lemma  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
222  | 
  shows open_halfspace_Re_lt: "open {z. Re(z) < b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
223  | 
    and open_halfspace_Re_gt: "open {z. Re(z) > b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
224  | 
    and closed_halfspace_Re_ge: "closed {z. Re(z) \<ge> b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
225  | 
    and closed_halfspace_Re_le: "closed {z. Re(z) \<le> b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
226  | 
    and closed_halfspace_Re_eq: "closed {z. Re(z) = b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
227  | 
    and open_halfspace_Im_lt: "open {z. Im(z) < b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
228  | 
    and open_halfspace_Im_gt: "open {z. Im(z) > b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
229  | 
    and closed_halfspace_Im_ge: "closed {z. Im(z) \<ge> b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
230  | 
    and closed_halfspace_Im_le: "closed {z. Im(z) \<le> b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
231  | 
    and closed_halfspace_Im_eq: "closed {z. Im(z) = b}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
232  | 
by (intro open_Collect_less closed_Collect_le closed_Collect_eq isCont_Re  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
233  | 
isCont_Im isCont_ident isCont_const)+  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
234  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
235  | 
lemma closed_complex_Reals: "closed (Reals :: complex set)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
236  | 
proof -  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
237  | 
  have "(Reals :: complex set) = {z. Im z = 0}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
238  | 
by (auto simp: complex_is_Real_iff)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
239  | 
then show ?thesis  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
240  | 
by (metis closed_halfspace_Im_eq)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
241  | 
qed  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
242  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
243  | 
lemma real_lim:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
244  | 
fixes l::complex  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
245  | 
assumes "(f ---> l) F" and "~(trivial_limit F)" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
246  | 
shows "l \<in> \<real>"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
247  | 
proof (rule Lim_in_closed_set[OF closed_complex_Reals _ assms(2,1)])  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
248  | 
show "eventually (\<lambda>x. f x \<in> \<real>) F"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
249  | 
using assms(3, 4) by (auto intro: eventually_mono)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
250  | 
qed  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
251  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
252  | 
lemma real_lim_sequentially:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
253  | 
fixes l::complex  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
254  | 
shows "(f ---> l) sequentially \<Longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
255  | 
by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
256  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
257  | 
lemma real_series:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
258  | 
fixes l::complex  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
259  | 
shows "f sums l \<Longrightarrow> (\<And>n. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
260  | 
unfolding sums_def  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
261  | 
by (metis real_lim_sequentially setsum_in_Reals)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
262  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
263  | 
lemma Lim_null_comparison_Re:  | 
| 
56889
 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 
hoelzl 
parents: 
56479 
diff
changeset
 | 
264  | 
assumes "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F" "(g ---> 0) F" shows "(f ---> 0) F"  | 
| 
 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 
hoelzl 
parents: 
56479 
diff
changeset
 | 
265  | 
by (rule Lim_null_comparison[OF assms(1)] tendsto_eq_intros assms(2))+ simp  | 
| 56215 | 266  | 
|
267  | 
subsection{*Holomorphic functions*}
 | 
|
268  | 
||
269  | 
definition complex_differentiable :: "[complex \<Rightarrow> complex, complex filter] \<Rightarrow> bool"  | 
|
270  | 
(infixr "(complex'_differentiable)" 50)  | 
|
271  | 
where "f complex_differentiable F \<equiv> \<exists>f'. (f has_field_derivative f') F"  | 
|
272  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
273  | 
lemma complex_differentiable_imp_continuous_at:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
274  | 
"f complex_differentiable (at x within s) \<Longrightarrow> continuous (at x within s) f"  | 
| 56215 | 275  | 
by (metis DERIV_continuous complex_differentiable_def)  | 
276  | 
||
277  | 
lemma complex_differentiable_within_subset:  | 
|
278  | 
"\<lbrakk>f complex_differentiable (at x within s); t \<subseteq> s\<rbrakk>  | 
|
279  | 
\<Longrightarrow> f complex_differentiable (at x within t)"  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
280  | 
by (metis DERIV_subset complex_differentiable_def)  | 
| 56215 | 281  | 
|
282  | 
lemma complex_differentiable_at_within:  | 
|
283  | 
"\<lbrakk>f complex_differentiable (at x)\<rbrakk>  | 
|
284  | 
\<Longrightarrow> f complex_differentiable (at x within s)"  | 
|
285  | 
unfolding complex_differentiable_def  | 
|
286  | 
by (metis DERIV_subset top_greatest)  | 
|
287  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
288  | 
lemma complex_differentiable_linear: "(op * c) complex_differentiable F"  | 
| 56215 | 289  | 
proof -  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
290  | 
show ?thesis  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
291  | 
unfolding complex_differentiable_def has_field_derivative_def mult_commute_abs  | 
| 56215 | 292  | 
by (force intro: has_derivative_mult_right)  | 
293  | 
qed  | 
|
294  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
295  | 
lemma complex_differentiable_const: "(\<lambda>z. c) complex_differentiable F"  | 
| 56215 | 296  | 
unfolding complex_differentiable_def has_field_derivative_def  | 
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
297  | 
by (rule exI [where x=0])  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
298  | 
(metis has_derivative_const lambda_zero)  | 
| 56215 | 299  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
300  | 
lemma complex_differentiable_ident: "(\<lambda>z. z) complex_differentiable F"  | 
| 56215 | 301  | 
unfolding complex_differentiable_def has_field_derivative_def  | 
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
302  | 
by (rule exI [where x=1])  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
303  | 
(simp add: lambda_one [symmetric])  | 
| 56215 | 304  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
305  | 
lemma complex_differentiable_id: "id complex_differentiable F"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
306  | 
unfolding id_def by (rule complex_differentiable_ident)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
307  | 
|
| 56215 | 308  | 
lemma complex_differentiable_minus:  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
309  | 
"f complex_differentiable F \<Longrightarrow> (\<lambda>z. - (f z)) complex_differentiable F"  | 
| 56215 | 310  | 
using assms unfolding complex_differentiable_def  | 
311  | 
by (metis field_differentiable_minus)  | 
|
312  | 
||
313  | 
lemma complex_differentiable_add:  | 
|
314  | 
assumes "f complex_differentiable F" "g complex_differentiable F"  | 
|
315  | 
shows "(\<lambda>z. f z + g z) complex_differentiable F"  | 
|
316  | 
using assms unfolding complex_differentiable_def  | 
|
317  | 
by (metis field_differentiable_add)  | 
|
318  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
319  | 
lemma complex_differentiable_setsum:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
320  | 
"(\<And>i. i \<in> I \<Longrightarrow> (f i) complex_differentiable F) \<Longrightarrow> (\<lambda>z. \<Sum>i\<in>I. f i z) complex_differentiable F"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
321  | 
by (induct I rule: infinite_finite_induct)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
322  | 
(auto intro: complex_differentiable_add complex_differentiable_const)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
323  | 
|
| 56215 | 324  | 
lemma complex_differentiable_diff:  | 
325  | 
assumes "f complex_differentiable F" "g complex_differentiable F"  | 
|
326  | 
shows "(\<lambda>z. f z - g z) complex_differentiable F"  | 
|
327  | 
using assms unfolding complex_differentiable_def  | 
|
328  | 
by (metis field_differentiable_diff)  | 
|
329  | 
||
330  | 
lemma complex_differentiable_inverse:  | 
|
331  | 
assumes "f complex_differentiable (at a within s)" "f a \<noteq> 0"  | 
|
332  | 
shows "(\<lambda>z. inverse (f z)) complex_differentiable (at a within s)"  | 
|
333  | 
using assms unfolding complex_differentiable_def  | 
|
334  | 
by (metis DERIV_inverse_fun)  | 
|
335  | 
||
336  | 
lemma complex_differentiable_mult:  | 
|
337  | 
assumes "f complex_differentiable (at a within s)"  | 
|
338  | 
"g complex_differentiable (at a within s)"  | 
|
339  | 
shows "(\<lambda>z. f z * g z) complex_differentiable (at a within s)"  | 
|
340  | 
using assms unfolding complex_differentiable_def  | 
|
341  | 
by (metis DERIV_mult [of f _ a s g])  | 
|
342  | 
||
343  | 
lemma complex_differentiable_divide:  | 
|
344  | 
assumes "f complex_differentiable (at a within s)"  | 
|
345  | 
"g complex_differentiable (at a within s)"  | 
|
346  | 
"g a \<noteq> 0"  | 
|
347  | 
shows "(\<lambda>z. f z / g z) complex_differentiable (at a within s)"  | 
|
348  | 
using assms unfolding complex_differentiable_def  | 
|
349  | 
by (metis DERIV_divide [of f _ a s g])  | 
|
350  | 
||
351  | 
lemma complex_differentiable_power:  | 
|
352  | 
assumes "f complex_differentiable (at a within s)"  | 
|
353  | 
shows "(\<lambda>z. f z ^ n) complex_differentiable (at a within s)"  | 
|
354  | 
using assms unfolding complex_differentiable_def  | 
|
355  | 
by (metis DERIV_power)  | 
|
356  | 
||
357  | 
lemma complex_differentiable_transform_within:  | 
|
358  | 
"0 < d \<Longrightarrow>  | 
|
359  | 
x \<in> s \<Longrightarrow>  | 
|
360  | 
(\<And>x'. x' \<in> s \<Longrightarrow> dist x' x < d \<Longrightarrow> f x' = g x') \<Longrightarrow>  | 
|
361  | 
f complex_differentiable (at x within s)  | 
|
362  | 
\<Longrightarrow> g complex_differentiable (at x within s)"  | 
|
363  | 
unfolding complex_differentiable_def has_field_derivative_def  | 
|
364  | 
by (blast intro: has_derivative_transform_within)  | 
|
365  | 
||
366  | 
lemma complex_differentiable_compose_within:  | 
|
367  | 
assumes "f complex_differentiable (at a within s)"  | 
|
368  | 
"g complex_differentiable (at (f a) within f`s)"  | 
|
369  | 
shows "(g o f) complex_differentiable (at a within s)"  | 
|
370  | 
using assms unfolding complex_differentiable_def  | 
|
371  | 
by (metis DERIV_image_chain)  | 
|
372  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
373  | 
lemma complex_differentiable_compose:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
374  | 
"f complex_differentiable at z \<Longrightarrow> g complex_differentiable at (f z)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
375  | 
\<Longrightarrow> (g o f) complex_differentiable at z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
376  | 
by (metis complex_differentiable_at_within complex_differentiable_compose_within)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
377  | 
|
| 56215 | 378  | 
lemma complex_differentiable_within_open:  | 
379  | 
"\<lbrakk>a \<in> s; open s\<rbrakk> \<Longrightarrow> f complex_differentiable at a within s \<longleftrightarrow>  | 
|
380  | 
f complex_differentiable at a"  | 
|
381  | 
unfolding complex_differentiable_def  | 
|
382  | 
by (metis at_within_open)  | 
|
383  | 
||
384  | 
subsection{*Caratheodory characterization.*}
 | 
|
385  | 
||
386  | 
lemma complex_differentiable_caratheodory_at:  | 
|
387  | 
"f complex_differentiable (at z) \<longleftrightarrow>  | 
|
388  | 
(\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z) g)"  | 
|
389  | 
using CARAT_DERIV [of f]  | 
|
390  | 
by (simp add: complex_differentiable_def has_field_derivative_def)  | 
|
391  | 
||
392  | 
lemma complex_differentiable_caratheodory_within:  | 
|
393  | 
"f complex_differentiable (at z within s) \<longleftrightarrow>  | 
|
394  | 
(\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z within s) g)"  | 
|
395  | 
using DERIV_caratheodory_within [of f]  | 
|
396  | 
by (simp add: complex_differentiable_def has_field_derivative_def)  | 
|
397  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
398  | 
subsection{*Holomorphic*}
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
399  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
400  | 
definition holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
401  | 
(infixl "(holomorphic'_on)" 50)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
402  | 
where "f holomorphic_on s \<equiv> \<forall>x\<in>s. f complex_differentiable (at x within s)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
403  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
404  | 
lemma holomorphic_on_empty: "f holomorphic_on {}"
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
405  | 
by (simp add: holomorphic_on_def)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
406  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
407  | 
lemma holomorphic_on_open:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
408  | 
"open s \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>f'. DERIV f x :> f')"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
409  | 
by (auto simp: holomorphic_on_def complex_differentiable_def has_field_derivative_def at_within_open [of _ s])  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
410  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
411  | 
lemma holomorphic_on_imp_continuous_on:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
412  | 
"f holomorphic_on s \<Longrightarrow> continuous_on s f"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
413  | 
by (metis complex_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
414  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
415  | 
lemma holomorphic_on_subset:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
416  | 
"f holomorphic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f holomorphic_on t"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
417  | 
unfolding holomorphic_on_def  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
418  | 
by (metis complex_differentiable_within_subset subsetD)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
419  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
420  | 
lemma holomorphic_transform: "\<lbrakk>f holomorphic_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
421  | 
by (metis complex_differentiable_transform_within linordered_field_no_ub holomorphic_on_def)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
422  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
423  | 
lemma holomorphic_cong: "s = t ==> (\<And>x. x \<in> s \<Longrightarrow> f x = g x) \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> g holomorphic_on t"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
424  | 
by (metis holomorphic_transform)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
425  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
426  | 
lemma holomorphic_on_linear: "(op * c) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
427  | 
unfolding holomorphic_on_def by (metis complex_differentiable_linear)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
428  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
429  | 
lemma holomorphic_on_const: "(\<lambda>z. c) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
430  | 
unfolding holomorphic_on_def by (metis complex_differentiable_const)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
431  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
432  | 
lemma holomorphic_on_ident: "(\<lambda>x. x) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
433  | 
unfolding holomorphic_on_def by (metis complex_differentiable_ident)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
434  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
435  | 
lemma holomorphic_on_id: "id holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
436  | 
unfolding id_def by (rule holomorphic_on_ident)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
437  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
438  | 
lemma holomorphic_on_compose:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
439  | 
"f holomorphic_on s \<Longrightarrow> g holomorphic_on (f ` s) \<Longrightarrow> (g o f) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
440  | 
using complex_differentiable_compose_within[of f _ s g]  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
441  | 
by (auto simp: holomorphic_on_def)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
442  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
443  | 
lemma holomorphic_on_compose_gen:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
444  | 
"f holomorphic_on s \<Longrightarrow> g holomorphic_on t \<Longrightarrow> f ` s \<subseteq> t \<Longrightarrow> (g o f) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
445  | 
by (metis holomorphic_on_compose holomorphic_on_subset)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
446  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
447  | 
lemma holomorphic_on_minus: "f holomorphic_on s \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
448  | 
by (metis complex_differentiable_minus holomorphic_on_def)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
449  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
450  | 
lemma holomorphic_on_add:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
451  | 
"\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
452  | 
unfolding holomorphic_on_def by (metis complex_differentiable_add)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
453  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
454  | 
lemma holomorphic_on_diff:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
455  | 
"\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z - g z) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
456  | 
unfolding holomorphic_on_def by (metis complex_differentiable_diff)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
457  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
458  | 
lemma holomorphic_on_mult:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
459  | 
"\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z * g z) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
460  | 
unfolding holomorphic_on_def by (metis complex_differentiable_mult)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
461  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
462  | 
lemma holomorphic_on_inverse:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
463  | 
"\<lbrakk>f holomorphic_on s; \<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. inverse (f z)) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
464  | 
unfolding holomorphic_on_def by (metis complex_differentiable_inverse)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
465  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
466  | 
lemma holomorphic_on_divide:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
467  | 
"\<lbrakk>f holomorphic_on s; g holomorphic_on s; \<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0\<rbrakk> \<Longrightarrow> (\<lambda>z. f z / g z) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
468  | 
unfolding holomorphic_on_def by (metis complex_differentiable_divide)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
469  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
470  | 
lemma holomorphic_on_power:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
471  | 
"f holomorphic_on s \<Longrightarrow> (\<lambda>z. (f z)^n) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
472  | 
unfolding holomorphic_on_def by (metis complex_differentiable_power)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
473  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
474  | 
lemma holomorphic_on_setsum:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
475  | 
"(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
476  | 
unfolding holomorphic_on_def by (metis complex_differentiable_setsum)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
477  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
478  | 
definition deriv :: "('a \<Rightarrow> 'a::real_normed_field) \<Rightarrow> 'a \<Rightarrow> 'a" where
 | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
479  | 
"deriv f x \<equiv> THE D. DERIV f x :> D"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
480  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
481  | 
lemma DERIV_imp_deriv: "DERIV f x :> f' \<Longrightarrow> deriv f x = f'"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
482  | 
unfolding deriv_def by (metis the_equality DERIV_unique)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
483  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
484  | 
lemma DERIV_deriv_iff_real_differentiable:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
485  | 
fixes x :: real  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
486  | 
shows "DERIV f x :> deriv f x \<longleftrightarrow> f differentiable at x"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
487  | 
unfolding differentiable_def by (metis DERIV_imp_deriv has_real_derivative_iff)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
488  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
489  | 
lemma real_derivative_chain:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
490  | 
fixes x :: real  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
491  | 
shows "f differentiable at x \<Longrightarrow> g differentiable at (f x)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
492  | 
\<Longrightarrow> deriv (g o f) x = deriv g (f x) * deriv f x"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
493  | 
by (metis DERIV_deriv_iff_real_differentiable DERIV_chain DERIV_imp_deriv)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
494  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
495  | 
lemma DERIV_deriv_iff_complex_differentiable:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
496  | 
"DERIV f x :> deriv f x \<longleftrightarrow> f complex_differentiable at x"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
497  | 
unfolding complex_differentiable_def by (metis DERIV_imp_deriv)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
498  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
499  | 
lemma complex_derivative_chain:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
500  | 
"f complex_differentiable at x \<Longrightarrow> g complex_differentiable at (f x)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
501  | 
\<Longrightarrow> deriv (g o f) x = deriv g (f x) * deriv f x"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
502  | 
by (metis DERIV_deriv_iff_complex_differentiable DERIV_chain DERIV_imp_deriv)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
503  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
504  | 
lemma complex_derivative_linear: "deriv (\<lambda>w. c * w) = (\<lambda>z. c)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
505  | 
by (metis DERIV_imp_deriv DERIV_cmult_Id)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
506  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
507  | 
lemma complex_derivative_ident: "deriv (\<lambda>w. w) = (\<lambda>z. 1)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
508  | 
by (metis DERIV_imp_deriv DERIV_ident)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
509  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
510  | 
lemma complex_derivative_const: "deriv (\<lambda>w. c) = (\<lambda>z. 0)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
511  | 
by (metis DERIV_imp_deriv DERIV_const)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
512  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
513  | 
lemma complex_derivative_add:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
514  | 
"\<lbrakk>f complex_differentiable at z; g complex_differentiable at z\<rbrakk>  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
515  | 
\<Longrightarrow> deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
516  | 
unfolding DERIV_deriv_iff_complex_differentiable[symmetric]  | 
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
517  | 
by (auto intro!: DERIV_imp_deriv derivative_intros)  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
518  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
519  | 
lemma complex_derivative_diff:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
520  | 
"\<lbrakk>f complex_differentiable at z; g complex_differentiable at z\<rbrakk>  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
521  | 
\<Longrightarrow> deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
522  | 
unfolding DERIV_deriv_iff_complex_differentiable[symmetric]  | 
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
523  | 
by (auto intro!: DERIV_imp_deriv derivative_intros)  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
524  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
525  | 
lemma complex_derivative_mult:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
526  | 
"\<lbrakk>f complex_differentiable at z; g complex_differentiable at z\<rbrakk>  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
527  | 
\<Longrightarrow> deriv (\<lambda>w. f w * g w) z = f z * deriv g z + deriv f z * g z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
528  | 
unfolding DERIV_deriv_iff_complex_differentiable[symmetric]  | 
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
529  | 
by (auto intro!: DERIV_imp_deriv derivative_eq_intros)  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
530  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
531  | 
lemma complex_derivative_cmult:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
532  | 
"f complex_differentiable at z \<Longrightarrow> deriv (\<lambda>w. c * f w) z = c * deriv f z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
533  | 
unfolding DERIV_deriv_iff_complex_differentiable[symmetric]  | 
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
534  | 
by (auto intro!: DERIV_imp_deriv derivative_eq_intros)  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
535  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
536  | 
lemma complex_derivative_cmult_right:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
537  | 
"f complex_differentiable at z \<Longrightarrow> deriv (\<lambda>w. f w * c) z = deriv f z * c"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
538  | 
unfolding DERIV_deriv_iff_complex_differentiable[symmetric]  | 
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
539  | 
by (auto intro!: DERIV_imp_deriv derivative_eq_intros)  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
540  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
541  | 
lemma complex_derivative_transform_within_open:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
542  | 
"\<lbrakk>f holomorphic_on s; g holomorphic_on s; open s; z \<in> s; \<And>w. w \<in> s \<Longrightarrow> f w = g w\<rbrakk>  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
543  | 
\<Longrightarrow> deriv f z = deriv g z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
544  | 
unfolding holomorphic_on_def  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
545  | 
by (rule DERIV_imp_deriv)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
546  | 
(metis DERIV_deriv_iff_complex_differentiable DERIV_transform_within_open at_within_open)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
547  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
548  | 
lemma complex_derivative_compose_linear:  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
549  | 
"f complex_differentiable at (c * z) \<Longrightarrow> deriv (\<lambda>w. f (c * w)) z = c * deriv f (c * z)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
550  | 
apply (rule DERIV_imp_deriv)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
551  | 
apply (simp add: DERIV_deriv_iff_complex_differentiable [symmetric])  | 
| 
59554
 
4044f53326c9
inlined rules to free user-space from technical names
 
haftmann 
parents: 
58877 
diff
changeset
 | 
552  | 
apply (drule DERIV_chain' [of "times c" c z UNIV f "deriv f (c * z)", OF DERIV_cmult_Id])  | 
| 
 
4044f53326c9
inlined rules to free user-space from technical names
 
haftmann 
parents: 
58877 
diff
changeset
 | 
553  | 
apply (simp add: algebra_simps)  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
554  | 
done  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
555  | 
|
| 56215 | 556  | 
subsection{*analyticity on a set*}
 | 
557  | 
||
558  | 
definition analytic_on (infixl "(analytic'_on)" 50)  | 
|
559  | 
where  | 
|
560  | 
"f analytic_on s \<equiv> \<forall>x \<in> s. \<exists>e. 0 < e \<and> f holomorphic_on (ball x e)"  | 
|
561  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
562  | 
lemma analytic_imp_holomorphic: "f analytic_on s \<Longrightarrow> f holomorphic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
563  | 
by (simp add: at_within_open [OF _ open_ball] analytic_on_def holomorphic_on_def)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
564  | 
(metis centre_in_ball complex_differentiable_at_within)  | 
| 56215 | 565  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
566  | 
lemma analytic_on_open: "open s \<Longrightarrow> f analytic_on s \<longleftrightarrow> f holomorphic_on s"  | 
| 56215 | 567  | 
apply (auto simp: analytic_imp_holomorphic)  | 
568  | 
apply (auto simp: analytic_on_def holomorphic_on_def)  | 
|
569  | 
by (metis holomorphic_on_def holomorphic_on_subset open_contains_ball)  | 
|
570  | 
||
571  | 
lemma analytic_on_imp_differentiable_at:  | 
|
572  | 
"f analytic_on s \<Longrightarrow> x \<in> s \<Longrightarrow> f complex_differentiable (at x)"  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
573  | 
apply (auto simp: analytic_on_def holomorphic_on_def)  | 
| 56215 | 574  | 
by (metis Topology_Euclidean_Space.open_ball centre_in_ball complex_differentiable_within_open)  | 
575  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
576  | 
lemma analytic_on_subset: "f analytic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f analytic_on t"  | 
| 56215 | 577  | 
by (auto simp: analytic_on_def)  | 
578  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
579  | 
lemma analytic_on_Un: "f analytic_on (s \<union> t) \<longleftrightarrow> f analytic_on s \<and> f analytic_on t"  | 
| 56215 | 580  | 
by (auto simp: analytic_on_def)  | 
581  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
582  | 
lemma analytic_on_Union: "f analytic_on (\<Union> s) \<longleftrightarrow> (\<forall>t \<in> s. f analytic_on t)"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
583  | 
by (auto simp: analytic_on_def)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
584  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
585  | 
lemma analytic_on_UN: "f analytic_on (\<Union>i\<in>I. s i) \<longleftrightarrow> (\<forall>i\<in>I. f analytic_on (s i))"  | 
| 56215 | 586  | 
by (auto simp: analytic_on_def)  | 
587  | 
||
588  | 
lemma analytic_on_holomorphic:  | 
|
589  | 
"f analytic_on s \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f holomorphic_on t)"  | 
|
590  | 
(is "?lhs = ?rhs")  | 
|
591  | 
proof -  | 
|
592  | 
have "?lhs \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t)"  | 
|
593  | 
proof safe  | 
|
594  | 
assume "f analytic_on s"  | 
|
595  | 
then show "\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t"  | 
|
596  | 
apply (simp add: analytic_on_def)  | 
|
597  | 
      apply (rule exI [where x="\<Union>{u. open u \<and> f analytic_on u}"], auto)
 | 
|
598  | 
apply (metis Topology_Euclidean_Space.open_ball analytic_on_open centre_in_ball)  | 
|
599  | 
by (metis analytic_on_def)  | 
|
600  | 
next  | 
|
601  | 
fix t  | 
|
602  | 
assume "open t" "s \<subseteq> t" "f analytic_on t"  | 
|
603  | 
then show "f analytic_on s"  | 
|
604  | 
by (metis analytic_on_subset)  | 
|
605  | 
qed  | 
|
606  | 
also have "... \<longleftrightarrow> ?rhs"  | 
|
607  | 
by (auto simp: analytic_on_open)  | 
|
608  | 
finally show ?thesis .  | 
|
609  | 
qed  | 
|
610  | 
||
611  | 
lemma analytic_on_linear: "(op * c) analytic_on s"  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
612  | 
by (auto simp add: analytic_on_holomorphic holomorphic_on_linear)  | 
| 56215 | 613  | 
|
614  | 
lemma analytic_on_const: "(\<lambda>z. c) analytic_on s"  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
615  | 
by (metis analytic_on_def holomorphic_on_const zero_less_one)  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
616  | 
|
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
617  | 
lemma analytic_on_ident: "(\<lambda>x. x) analytic_on s"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
618  | 
by (simp add: analytic_on_def holomorphic_on_ident gt_ex)  | 
| 56215 | 619  | 
|
620  | 
lemma analytic_on_id: "id analytic_on s"  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
621  | 
unfolding id_def by (rule analytic_on_ident)  | 
| 56215 | 622  | 
|
623  | 
lemma analytic_on_compose:  | 
|
624  | 
assumes f: "f analytic_on s"  | 
|
625  | 
and g: "g analytic_on (f ` s)"  | 
|
626  | 
shows "(g o f) analytic_on s"  | 
|
627  | 
unfolding analytic_on_def  | 
|
628  | 
proof (intro ballI)  | 
|
629  | 
fix x  | 
|
630  | 
assume x: "x \<in> s"  | 
|
631  | 
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f  | 
|
632  | 
by (metis analytic_on_def)  | 
|
633  | 
obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g  | 
|
634  | 
by (metis analytic_on_def g image_eqI x)  | 
|
635  | 
have "isCont f x"  | 
|
636  | 
by (metis analytic_on_imp_differentiable_at complex_differentiable_imp_continuous_at f x)  | 
|
637  | 
with e' obtain d where d: "0 < d" and fd: "f ` ball x d \<subseteq> ball (f x) e'"  | 
|
638  | 
by (auto simp: continuous_at_ball)  | 
|
639  | 
have "g \<circ> f holomorphic_on ball x (min d e)"  | 
|
640  | 
apply (rule holomorphic_on_compose)  | 
|
641  | 
apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)  | 
|
642  | 
by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball)  | 
|
643  | 
then show "\<exists>e>0. g \<circ> f holomorphic_on ball x e"  | 
|
644  | 
by (metis d e min_less_iff_conj)  | 
|
645  | 
qed  | 
|
646  | 
||
647  | 
lemma analytic_on_compose_gen:  | 
|
648  | 
"f analytic_on s \<Longrightarrow> g analytic_on t \<Longrightarrow> (\<And>z. z \<in> s \<Longrightarrow> f z \<in> t)  | 
|
649  | 
\<Longrightarrow> g o f analytic_on s"  | 
|
650  | 
by (metis analytic_on_compose analytic_on_subset image_subset_iff)  | 
|
651  | 
||
652  | 
lemma analytic_on_neg:  | 
|
653  | 
"f analytic_on s \<Longrightarrow> (\<lambda>z. -(f z)) analytic_on s"  | 
|
654  | 
by (metis analytic_on_holomorphic holomorphic_on_minus)  | 
|
655  | 
||
656  | 
lemma analytic_on_add:  | 
|
657  | 
assumes f: "f analytic_on s"  | 
|
658  | 
and g: "g analytic_on s"  | 
|
659  | 
shows "(\<lambda>z. f z + g z) analytic_on s"  | 
|
660  | 
unfolding analytic_on_def  | 
|
661  | 
proof (intro ballI)  | 
|
662  | 
fix z  | 
|
663  | 
assume z: "z \<in> s"  | 
|
664  | 
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f  | 
|
665  | 
by (metis analytic_on_def)  | 
|
666  | 
obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g  | 
|
667  | 
by (metis analytic_on_def g z)  | 
|
668  | 
have "(\<lambda>z. f z + g z) holomorphic_on ball z (min e e')"  | 
|
669  | 
apply (rule holomorphic_on_add)  | 
|
670  | 
apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)  | 
|
671  | 
by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)  | 
|
672  | 
then show "\<exists>e>0. (\<lambda>z. f z + g z) holomorphic_on ball z e"  | 
|
673  | 
by (metis e e' min_less_iff_conj)  | 
|
674  | 
qed  | 
|
675  | 
||
676  | 
lemma analytic_on_diff:  | 
|
677  | 
assumes f: "f analytic_on s"  | 
|
678  | 
and g: "g analytic_on s"  | 
|
679  | 
shows "(\<lambda>z. f z - g z) analytic_on s"  | 
|
680  | 
unfolding analytic_on_def  | 
|
681  | 
proof (intro ballI)  | 
|
682  | 
fix z  | 
|
683  | 
assume z: "z \<in> s"  | 
|
684  | 
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f  | 
|
685  | 
by (metis analytic_on_def)  | 
|
686  | 
obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g  | 
|
687  | 
by (metis analytic_on_def g z)  | 
|
688  | 
have "(\<lambda>z. f z - g z) holomorphic_on ball z (min e e')"  | 
|
689  | 
apply (rule holomorphic_on_diff)  | 
|
690  | 
apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)  | 
|
691  | 
by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)  | 
|
692  | 
then show "\<exists>e>0. (\<lambda>z. f z - g z) holomorphic_on ball z e"  | 
|
693  | 
by (metis e e' min_less_iff_conj)  | 
|
694  | 
qed  | 
|
695  | 
||
696  | 
lemma analytic_on_mult:  | 
|
697  | 
assumes f: "f analytic_on s"  | 
|
698  | 
and g: "g analytic_on s"  | 
|
699  | 
shows "(\<lambda>z. f z * g z) analytic_on s"  | 
|
700  | 
unfolding analytic_on_def  | 
|
701  | 
proof (intro ballI)  | 
|
702  | 
fix z  | 
|
703  | 
assume z: "z \<in> s"  | 
|
704  | 
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f  | 
|
705  | 
by (metis analytic_on_def)  | 
|
706  | 
obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g  | 
|
707  | 
by (metis analytic_on_def g z)  | 
|
708  | 
have "(\<lambda>z. f z * g z) holomorphic_on ball z (min e e')"  | 
|
709  | 
apply (rule holomorphic_on_mult)  | 
|
710  | 
apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)  | 
|
711  | 
by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)  | 
|
712  | 
then show "\<exists>e>0. (\<lambda>z. f z * g z) holomorphic_on ball z e"  | 
|
713  | 
by (metis e e' min_less_iff_conj)  | 
|
714  | 
qed  | 
|
715  | 
||
716  | 
lemma analytic_on_inverse:  | 
|
717  | 
assumes f: "f analytic_on s"  | 
|
718  | 
and nz: "(\<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0)"  | 
|
719  | 
shows "(\<lambda>z. inverse (f z)) analytic_on s"  | 
|
720  | 
unfolding analytic_on_def  | 
|
721  | 
proof (intro ballI)  | 
|
722  | 
fix z  | 
|
723  | 
assume z: "z \<in> s"  | 
|
724  | 
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f  | 
|
725  | 
by (metis analytic_on_def)  | 
|
726  | 
have "continuous_on (ball z e) f"  | 
|
727  | 
by (metis fh holomorphic_on_imp_continuous_on)  | 
|
728  | 
then obtain e' where e': "0 < e'" and nz': "\<And>y. dist z y < e' \<Longrightarrow> f y \<noteq> 0"  | 
|
729  | 
by (metis Topology_Euclidean_Space.open_ball centre_in_ball continuous_on_open_avoid e z nz)  | 
|
730  | 
have "(\<lambda>z. inverse (f z)) holomorphic_on ball z (min e e')"  | 
|
731  | 
apply (rule holomorphic_on_inverse)  | 
|
732  | 
apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball)  | 
|
733  | 
by (metis nz' mem_ball min_less_iff_conj)  | 
|
734  | 
then show "\<exists>e>0. (\<lambda>z. inverse (f z)) holomorphic_on ball z e"  | 
|
735  | 
by (metis e e' min_less_iff_conj)  | 
|
736  | 
qed  | 
|
737  | 
||
738  | 
||
739  | 
lemma analytic_on_divide:  | 
|
740  | 
assumes f: "f analytic_on s"  | 
|
741  | 
and g: "g analytic_on s"  | 
|
742  | 
and nz: "(\<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0)"  | 
|
743  | 
shows "(\<lambda>z. f z / g z) analytic_on s"  | 
|
744  | 
unfolding divide_inverse  | 
|
745  | 
by (metis analytic_on_inverse analytic_on_mult f g nz)  | 
|
746  | 
||
747  | 
lemma analytic_on_power:  | 
|
748  | 
"f analytic_on s \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on s"  | 
|
749  | 
by (induct n) (auto simp: analytic_on_const analytic_on_mult)  | 
|
750  | 
||
751  | 
lemma analytic_on_setsum:  | 
|
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
752  | 
"(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on s) \<Longrightarrow> (\<lambda>x. setsum (\<lambda>i. f i x) I) analytic_on s"  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
753  | 
by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add)  | 
| 56215 | 754  | 
|
755  | 
subsection{*analyticity at a point.*}
 | 
|
756  | 
||
757  | 
lemma analytic_at_ball:  | 
|
758  | 
  "f analytic_on {z} \<longleftrightarrow> (\<exists>e. 0<e \<and> f holomorphic_on ball z e)"
 | 
|
759  | 
by (metis analytic_on_def singleton_iff)  | 
|
760  | 
||
761  | 
lemma analytic_at:  | 
|
762  | 
    "f analytic_on {z} \<longleftrightarrow> (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s)"
 | 
|
763  | 
by (metis analytic_on_holomorphic empty_subsetI insert_subset)  | 
|
764  | 
||
765  | 
lemma analytic_on_analytic_at:  | 
|
766  | 
    "f analytic_on s \<longleftrightarrow> (\<forall>z \<in> s. f analytic_on {z})"
 | 
|
767  | 
by (metis analytic_at_ball analytic_on_def)  | 
|
768  | 
||
769  | 
lemma analytic_at_two:  | 
|
770  | 
  "f analytic_on {z} \<and> g analytic_on {z} \<longleftrightarrow>
 | 
|
771  | 
(\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s \<and> g holomorphic_on s)"  | 
|
772  | 
(is "?lhs = ?rhs")  | 
|
773  | 
proof  | 
|
774  | 
assume ?lhs  | 
|
775  | 
then obtain s t  | 
|
776  | 
where st: "open s" "z \<in> s" "f holomorphic_on s"  | 
|
777  | 
"open t" "z \<in> t" "g holomorphic_on t"  | 
|
778  | 
by (auto simp: analytic_at)  | 
|
779  | 
show ?rhs  | 
|
780  | 
apply (rule_tac x="s \<inter> t" in exI)  | 
|
781  | 
using st  | 
|
782  | 
apply (auto simp: Diff_subset holomorphic_on_subset)  | 
|
783  | 
done  | 
|
784  | 
next  | 
|
785  | 
assume ?rhs  | 
|
786  | 
then show ?lhs  | 
|
787  | 
by (force simp add: analytic_at)  | 
|
788  | 
qed  | 
|
789  | 
||
790  | 
subsection{*Combining theorems for derivative with ``analytic at'' hypotheses*}
 | 
|
791  | 
||
792  | 
lemma  | 
|
793  | 
  assumes "f analytic_on {z}" "g analytic_on {z}"
 | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
794  | 
shows complex_derivative_add_at: "deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
795  | 
and complex_derivative_diff_at: "deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
796  | 
and complex_derivative_mult_at: "deriv (\<lambda>w. f w * g w) z =  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
797  | 
f z * deriv g z + deriv f z * g z"  | 
| 56215 | 798  | 
proof -  | 
799  | 
obtain s where s: "open s" "z \<in> s" "f holomorphic_on s" "g holomorphic_on s"  | 
|
800  | 
using assms by (metis analytic_at_two)  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
801  | 
show "deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
802  | 
apply (rule DERIV_imp_deriv [OF DERIV_add])  | 
| 56215 | 803  | 
using s  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
804  | 
apply (auto simp: holomorphic_on_open complex_differentiable_def DERIV_deriv_iff_complex_differentiable)  | 
| 56215 | 805  | 
done  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
806  | 
show "deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
807  | 
apply (rule DERIV_imp_deriv [OF DERIV_diff])  | 
| 56215 | 808  | 
using s  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
809  | 
apply (auto simp: holomorphic_on_open complex_differentiable_def DERIV_deriv_iff_complex_differentiable)  | 
| 56215 | 810  | 
done  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
811  | 
show "deriv (\<lambda>w. f w * g w) z = f z * deriv g z + deriv f z * g z"  | 
| 
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
812  | 
apply (rule DERIV_imp_deriv [OF DERIV_mult'])  | 
| 56215 | 813  | 
using s  | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
814  | 
apply (auto simp: holomorphic_on_open complex_differentiable_def DERIV_deriv_iff_complex_differentiable)  | 
| 56215 | 815  | 
done  | 
816  | 
qed  | 
|
817  | 
||
818  | 
lemma complex_derivative_cmult_at:  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
819  | 
  "f analytic_on {z} \<Longrightarrow>  deriv (\<lambda>w. c * f w) z = c * deriv f z"
 | 
| 56215 | 820  | 
by (auto simp: complex_derivative_mult_at complex_derivative_const analytic_on_const)  | 
821  | 
||
822  | 
lemma complex_derivative_cmult_right_at:  | 
|
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
823  | 
  "f analytic_on {z} \<Longrightarrow>  deriv (\<lambda>w. f w * c) z = deriv f z * c"
 | 
| 56215 | 824  | 
by (auto simp: complex_derivative_mult_at complex_derivative_const analytic_on_const)  | 
825  | 
||
826  | 
subsection{*Complex differentiation of sequences and series*}
 | 
|
827  | 
||
828  | 
lemma has_complex_derivative_sequence:  | 
|
829  | 
fixes s :: "complex set"  | 
|
830  | 
assumes cvs: "convex s"  | 
|
831  | 
and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)"  | 
|
832  | 
and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s \<longrightarrow> norm (f' n x - g' x) \<le> e"  | 
|
833  | 
and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) ---> l) sequentially"  | 
|
834  | 
shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) ---> g x) sequentially \<and>  | 
|
835  | 
(g has_field_derivative (g' x)) (at x within s)"  | 
|
836  | 
proof -  | 
|
837  | 
from assms obtain x l where x: "x \<in> s" and tf: "((\<lambda>n. f n x) ---> l) sequentially"  | 
|
838  | 
by blast  | 
|
839  | 
  { fix e::real assume e: "e > 0"
 | 
|
840  | 
then obtain N where N: "\<forall>n\<ge>N. \<forall>x. x \<in> s \<longrightarrow> cmod (f' n x - g' x) \<le> e"  | 
|
841  | 
by (metis conv)  | 
|
842  | 
have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"  | 
|
843  | 
proof (rule exI [of _ N], clarify)  | 
|
844  | 
fix n y h  | 
|
845  | 
assume "N \<le> n" "y \<in> s"  | 
|
846  | 
then have "cmod (f' n y - g' y) \<le> e"  | 
|
847  | 
by (metis N)  | 
|
848  | 
then have "cmod h * cmod (f' n y - g' y) \<le> cmod h * e"  | 
|
849  | 
by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)  | 
|
850  | 
then show "cmod (f' n y * h - g' y * h) \<le> e * cmod h"  | 
|
851  | 
by (simp add: norm_mult [symmetric] field_simps)  | 
|
852  | 
qed  | 
|
853  | 
} note ** = this  | 
|
854  | 
show ?thesis  | 
|
855  | 
unfolding has_field_derivative_def  | 
|
856  | 
proof (rule has_derivative_sequence [OF cvs _ _ x])  | 
|
857  | 
show "\<forall>n. \<forall>x\<in>s. (f n has_derivative (op * (f' n x))) (at x within s)"  | 
|
858  | 
by (metis has_field_derivative_def df)  | 
|
859  | 
next show "(\<lambda>n. f n x) ----> l"  | 
|
860  | 
by (rule tf)  | 
|
861  | 
next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod (f' n x * h - g' x * h) \<le> e * cmod h"  | 
|
862  | 
by (blast intro: **)  | 
|
863  | 
qed  | 
|
864  | 
qed  | 
|
865  | 
||
866  | 
||
867  | 
lemma has_complex_derivative_series:  | 
|
868  | 
fixes s :: "complex set"  | 
|
869  | 
assumes cvs: "convex s"  | 
|
870  | 
and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)"  | 
|
871  | 
and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s  | 
|
872  | 
\<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e"  | 
|
873  | 
and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) sums l)"  | 
|
874  | 
shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) sums g x) \<and> ((g has_field_derivative g' x) (at x within s))"  | 
|
875  | 
proof -  | 
|
876  | 
from assms obtain x l where x: "x \<in> s" and sf: "((\<lambda>n. f n x) sums l)"  | 
|
877  | 
by blast  | 
|
878  | 
  { fix e::real assume e: "e > 0"
 | 
|
879  | 
then obtain N where N: "\<forall>n x. n \<ge> N \<longrightarrow> x \<in> s  | 
|
880  | 
\<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e"  | 
|
881  | 
by (metis conv)  | 
|
882  | 
have "\<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h"  | 
|
883  | 
proof (rule exI [of _ N], clarify)  | 
|
884  | 
fix n y h  | 
|
885  | 
assume "N \<le> n" "y \<in> s"  | 
|
886  | 
then have "cmod ((\<Sum>i<n. f' i y) - g' y) \<le> e"  | 
|
887  | 
by (metis N)  | 
|
888  | 
then have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * e"  | 
|
889  | 
by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)  | 
|
890  | 
then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> e * cmod h"  | 
|
891  | 
by (simp add: norm_mult [symmetric] field_simps setsum_right_distrib)  | 
|
892  | 
qed  | 
|
893  | 
} note ** = this  | 
|
894  | 
show ?thesis  | 
|
895  | 
unfolding has_field_derivative_def  | 
|
896  | 
proof (rule has_derivative_series [OF cvs _ _ x])  | 
|
897  | 
fix n x  | 
|
898  | 
assume "x \<in> s"  | 
|
899  | 
then show "((f n) has_derivative (\<lambda>z. z * f' n x)) (at x within s)"  | 
|
900  | 
by (metis df has_field_derivative_def mult_commute_abs)  | 
|
901  | 
next show " ((\<lambda>n. f n x) sums l)"  | 
|
902  | 
by (rule sf)  | 
|
903  | 
next show "\<forall>e>0. \<exists>N. \<forall>n\<ge>N. \<forall>x\<in>s. \<forall>h. cmod ((\<Sum>i<n. h * f' i x) - g' x * h) \<le> e * cmod h"  | 
|
904  | 
by (blast intro: **)  | 
|
905  | 
qed  | 
|
906  | 
qed  | 
|
907  | 
||
908  | 
subsection{*Bound theorem*}
 | 
|
909  | 
||
910  | 
lemma complex_differentiable_bound:  | 
|
911  | 
fixes s :: "complex set"  | 
|
912  | 
assumes cvs: "convex s"  | 
|
913  | 
and df: "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z within s)"  | 
|
914  | 
and dn: "\<And>z. z \<in> s \<Longrightarrow> norm (f' z) \<le> B"  | 
|
915  | 
and "x \<in> s" "y \<in> s"  | 
|
916  | 
shows "norm(f x - f y) \<le> B * norm(x - y)"  | 
|
917  | 
apply (rule differentiable_bound [OF cvs])  | 
|
| 
56223
 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 
huffman 
parents: 
56217 
diff
changeset
 | 
918  | 
apply (rule ballI, erule df [unfolded has_field_derivative_def])  | 
| 
 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 
huffman 
parents: 
56217 
diff
changeset
 | 
919  | 
apply (rule ballI, rule onorm_le, simp add: norm_mult mult_right_mono dn)  | 
| 
 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 
huffman 
parents: 
56217 
diff
changeset
 | 
920  | 
apply fact  | 
| 
 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 
huffman 
parents: 
56217 
diff
changeset
 | 
921  | 
apply fact  | 
| 56215 | 922  | 
done  | 
923  | 
||
924  | 
subsection{*Inverse function theorem for complex derivatives.*}
 | 
|
925  | 
||
926  | 
lemma has_complex_derivative_inverse_basic:  | 
|
927  | 
fixes f :: "complex \<Rightarrow> complex"  | 
|
928  | 
shows "DERIV f (g y) :> f' \<Longrightarrow>  | 
|
929  | 
f' \<noteq> 0 \<Longrightarrow>  | 
|
930  | 
continuous (at y) g \<Longrightarrow>  | 
|
931  | 
open t \<Longrightarrow>  | 
|
932  | 
y \<in> t \<Longrightarrow>  | 
|
933  | 
(\<And>z. z \<in> t \<Longrightarrow> f (g z) = z)  | 
|
934  | 
\<Longrightarrow> DERIV g y :> inverse (f')"  | 
|
935  | 
unfolding has_field_derivative_def  | 
|
936  | 
apply (rule has_derivative_inverse_basic)  | 
|
937  | 
apply (auto simp: bounded_linear_mult_right)  | 
|
938  | 
done  | 
|
939  | 
||
940  | 
(*Used only once, in Multivariate/cauchy.ml. *)  | 
|
941  | 
lemma has_complex_derivative_inverse_strong:  | 
|
942  | 
fixes f :: "complex \<Rightarrow> complex"  | 
|
943  | 
shows "DERIV f x :> f' \<Longrightarrow>  | 
|
944  | 
f' \<noteq> 0 \<Longrightarrow>  | 
|
945  | 
open s \<Longrightarrow>  | 
|
946  | 
x \<in> s \<Longrightarrow>  | 
|
947  | 
continuous_on s f \<Longrightarrow>  | 
|
948  | 
(\<And>z. z \<in> s \<Longrightarrow> g (f z) = z)  | 
|
949  | 
\<Longrightarrow> DERIV g (f x) :> inverse (f')"  | 
|
950  | 
unfolding has_field_derivative_def  | 
|
951  | 
apply (rule has_derivative_inverse_strong [of s x f g ])  | 
|
952  | 
using assms  | 
|
953  | 
by auto  | 
|
954  | 
||
955  | 
lemma has_complex_derivative_inverse_strong_x:  | 
|
956  | 
fixes f :: "complex \<Rightarrow> complex"  | 
|
957  | 
shows "DERIV f (g y) :> f' \<Longrightarrow>  | 
|
958  | 
f' \<noteq> 0 \<Longrightarrow>  | 
|
959  | 
open s \<Longrightarrow>  | 
|
960  | 
continuous_on s f \<Longrightarrow>  | 
|
961  | 
g y \<in> s \<Longrightarrow> f(g y) = y \<Longrightarrow>  | 
|
962  | 
(\<And>z. z \<in> s \<Longrightarrow> g (f z) = z)  | 
|
963  | 
\<Longrightarrow> DERIV g y :> inverse (f')"  | 
|
964  | 
unfolding has_field_derivative_def  | 
|
965  | 
apply (rule has_derivative_inverse_strong_x [of s g y f])  | 
|
966  | 
using assms  | 
|
967  | 
by auto  | 
|
968  | 
||
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
969  | 
subsection {* Taylor on Complex Numbers *}
 | 
| 56215 | 970  | 
|
971  | 
lemma setsum_Suc_reindex:  | 
|
972  | 
fixes f :: "nat \<Rightarrow> 'a::ab_group_add"  | 
|
973  | 
    shows  "setsum f {0..n} = f 0 - f (Suc n) + setsum (\<lambda>i. f (Suc i)) {0..n}"
 | 
|
974  | 
by (induct n) auto  | 
|
975  | 
||
976  | 
lemma complex_taylor:  | 
|
977  | 
assumes s: "convex s"  | 
|
978  | 
and f: "\<And>i x. x \<in> s \<Longrightarrow> i \<le> n \<Longrightarrow> (f i has_field_derivative f (Suc i) x) (at x within s)"  | 
|
979  | 
and B: "\<And>x. x \<in> s \<Longrightarrow> cmod (f (Suc n) x) \<le> B"  | 
|
980  | 
and w: "w \<in> s"  | 
|
981  | 
and z: "z \<in> s"  | 
|
982  | 
shows "cmod(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / of_nat (fact i)))  | 
|
983  | 
\<le> B * cmod(z - w)^(Suc n) / fact n"  | 
|
984  | 
proof -  | 
|
985  | 
have wzs: "closed_segment w z \<subseteq> s" using assms  | 
|
986  | 
by (metis convex_contains_segment)  | 
|
987  | 
  { fix u
 | 
|
988  | 
assume "u \<in> closed_segment w z"  | 
|
989  | 
then have "u \<in> s"  | 
|
990  | 
by (metis wzs subsetD)  | 
|
991  | 
have "(\<Sum>i\<le>n. f i u * (- of_nat i * (z-u)^(i - 1)) / of_nat (fact i) +  | 
|
992  | 
f (Suc i) u * (z-u)^i / of_nat (fact i)) =  | 
|
993  | 
f (Suc n) u * (z-u) ^ n / of_nat (fact n)"  | 
|
994  | 
proof (induction n)  | 
|
995  | 
case 0 show ?case by simp  | 
|
996  | 
next  | 
|
997  | 
case (Suc n)  | 
|
998  | 
have "(\<Sum>i\<le>Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / of_nat (fact i) +  | 
|
999  | 
f (Suc i) u * (z-u) ^ i / of_nat (fact i)) =  | 
|
1000  | 
f (Suc n) u * (z-u) ^ n / of_nat (fact n) +  | 
|
1001  | 
f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / of_nat (fact (Suc n)) -  | 
|
1002  | 
f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / of_nat (fact (Suc n))"  | 
|
| 
56479
 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 
hoelzl 
parents: 
56409 
diff
changeset
 | 
1003  | 
using Suc by simp  | 
| 56215 | 1004  | 
also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / of_nat (fact (Suc n))"  | 
1005  | 
proof -  | 
|
1006  | 
have "of_nat(fact(Suc n)) *  | 
|
1007  | 
(f(Suc n) u *(z-u) ^ n / of_nat(fact n) +  | 
|
1008  | 
f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / of_nat(fact(Suc n)) -  | 
|
1009  | 
f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / of_nat(fact(Suc n))) =  | 
|
1010  | 
(of_nat(fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / of_nat(fact n) +  | 
|
1011  | 
(of_nat(fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / of_nat(fact(Suc n))) -  | 
|
1012  | 
(of_nat(fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / of_nat(fact(Suc n))"  | 
|
1013  | 
by (simp add: algebra_simps del: fact_Suc)  | 
|
1014  | 
also have "... =  | 
|
1015  | 
(of_nat (fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / of_nat (fact n) +  | 
|
1016  | 
(f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -  | 
|
1017  | 
(f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"  | 
|
1018  | 
by (simp del: fact_Suc)  | 
|
1019  | 
also have "... =  | 
|
1020  | 
(of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) +  | 
|
1021  | 
(f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -  | 
|
1022  | 
(f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"  | 
|
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
56889 
diff
changeset
 | 
1023  | 
by (simp only: fact_Suc of_nat_mult ac_simps) simp  | 
| 56215 | 1024  | 
also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)"  | 
1025  | 
by (simp add: algebra_simps)  | 
|
1026  | 
finally show ?thesis  | 
|
1027  | 
by (simp add: mult_left_cancel [where c = "of_nat (fact (Suc n))", THEN iffD1] del: fact_Suc)  | 
|
1028  | 
qed  | 
|
1029  | 
finally show ?case .  | 
|
1030  | 
qed  | 
|
1031  | 
then have "((\<lambda>v. (\<Sum>i\<le>n. f i v * (z - v)^i / of_nat (fact i)))  | 
|
1032  | 
has_field_derivative f (Suc n) u * (z-u) ^ n / of_nat (fact n))  | 
|
1033  | 
(at u within s)"  | 
|
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1034  | 
apply (intro derivative_eq_intros)  | 
| 56215 | 1035  | 
apply (blast intro: assms `u \<in> s`)  | 
1036  | 
apply (rule refl)+  | 
|
1037  | 
apply (auto simp: field_simps)  | 
|
1038  | 
done  | 
|
1039  | 
} note sum_deriv = this  | 
|
1040  | 
  { fix u
 | 
|
1041  | 
assume u: "u \<in> closed_segment w z"  | 
|
1042  | 
then have us: "u \<in> s"  | 
|
1043  | 
by (metis wzs subsetD)  | 
|
1044  | 
have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> cmod (f (Suc n) u) * cmod (u - z) ^ n"  | 
|
1045  | 
by (metis norm_minus_commute order_refl)  | 
|
1046  | 
also have "... \<le> cmod (f (Suc n) u) * cmod (z - w) ^ n"  | 
|
1047  | 
by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u])  | 
|
1048  | 
also have "... \<le> B * cmod (z - w) ^ n"  | 
|
1049  | 
by (metis norm_ge_zero zero_le_power mult_right_mono B [OF us])  | 
|
1050  | 
finally have "cmod (f (Suc n) u) * cmod (z - u) ^ n \<le> B * cmod (z - w) ^ n" .  | 
|
1051  | 
} note cmod_bound = this  | 
|
1052  | 
have "(\<Sum>i\<le>n. f i z * (z - z) ^ i / of_nat (fact i)) = (\<Sum>i\<le>n. (f i z / of_nat (fact i)) * 0 ^ i)"  | 
|
1053  | 
by simp  | 
|
1054  | 
also have "\<dots> = f 0 z / of_nat (fact 0)"  | 
|
1055  | 
by (subst setsum_zero_power) simp  | 
|
1056  | 
finally have "cmod (f 0 z - (\<Sum>i\<le>n. f i w * (z - w) ^ i / of_nat (fact i)))  | 
|
1057  | 
\<le> cmod ((\<Sum>i\<le>n. f i w * (z - w) ^ i / of_nat (fact i)) -  | 
|
1058  | 
(\<Sum>i\<le>n. f i z * (z - z) ^ i / of_nat (fact i)))"  | 
|
1059  | 
by (simp add: norm_minus_commute)  | 
|
1060  | 
also have "... \<le> B * cmod (z - w) ^ n / real_of_nat (fact n) * cmod (w - z)"  | 
|
1061  | 
apply (rule complex_differentiable_bound  | 
|
1062  | 
[where f' = "\<lambda>w. f (Suc n) w * (z - w)^n / of_nat(fact n)"  | 
|
1063  | 
and s = "closed_segment w z", OF convex_segment])  | 
|
1064  | 
apply (auto simp: ends_in_segment real_of_nat_def DERIV_subset [OF sum_deriv wzs]  | 
|
1065  | 
norm_divide norm_mult norm_power divide_le_cancel cmod_bound)  | 
|
1066  | 
done  | 
|
1067  | 
also have "... \<le> B * cmod (z - w) ^ Suc n / real (fact n)"  | 
|
1068  | 
by (simp add: algebra_simps norm_minus_commute real_of_nat_def)  | 
|
1069  | 
finally show ?thesis .  | 
|
1070  | 
qed  | 
|
1071  | 
||
| 
56238
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1072  | 
text{* Something more like the traditional MVT for real components.*}
 | 
| 
56370
 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 
hoelzl 
parents: 
56369 
diff
changeset
 | 
1073  | 
|
| 
56238
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1074  | 
lemma complex_mvt_line:  | 
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
1075  | 
assumes "\<And>u. u \<in> closed_segment w z \<Longrightarrow> (f has_field_derivative f'(u)) (at u)"  | 
| 
56238
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1076  | 
shows "\<exists>u. u \<in> open_segment w z \<and> Re(f z) - Re(f w) = Re(f'(u) * (z - w))"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1077  | 
proof -  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1078  | 
have twz: "\<And>t. (1 - t) *\<^sub>R w + t *\<^sub>R z = w + t *\<^sub>R (z - w)"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1079  | 
by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib)  | 
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1080  | 
note assms[unfolded has_field_derivative_def, derivative_intros]  | 
| 
56238
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1081  | 
show ?thesis  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1082  | 
apply (cut_tac mvt_simple  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1083  | 
[of 0 1 "Re o f o (\<lambda>t. (1 - t) *\<^sub>R w + t *\<^sub>R z)"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1084  | 
"\<lambda>u. Re o (\<lambda>h. f'((1 - u) *\<^sub>R w + u *\<^sub>R z) * h) o (\<lambda>t. t *\<^sub>R (z - w))"])  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1085  | 
apply auto  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1086  | 
apply (rule_tac x="(1 - x) *\<^sub>R w + x *\<^sub>R z" in exI)  | 
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
1087  | 
apply (auto simp add: open_segment_def twz) []  | 
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1088  | 
apply (intro derivative_eq_intros has_derivative_at_within)  | 
| 
56369
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
1089  | 
apply simp_all  | 
| 
 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 
hoelzl 
parents: 
56332 
diff
changeset
 | 
1090  | 
apply (simp add: fun_eq_iff real_vector.scale_right_diff_distrib)  | 
| 
56238
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1091  | 
apply (force simp add: twz closed_segment_def)  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1092  | 
done  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1093  | 
qed  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1094  | 
|
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1095  | 
lemma complex_taylor_mvt:  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1096  | 
assumes "\<And>i x. \<lbrakk>x \<in> closed_segment w z; i \<le> n\<rbrakk> \<Longrightarrow> ((f i) has_field_derivative f (Suc i) x) (at x)"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1097  | 
shows "\<exists>u. u \<in> closed_segment w z \<and>  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1098  | 
Re (f 0 z) =  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1099  | 
Re ((\<Sum>i = 0..n. f i w * (z - w) ^ i / of_nat (fact i)) +  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1100  | 
(f (Suc n) u * (z-u)^n / of_nat (fact n)) * (z - w))"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1101  | 
proof -  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1102  | 
  { fix u
 | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1103  | 
assume u: "u \<in> closed_segment w z"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1104  | 
have "(\<Sum>i = 0..n.  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1105  | 
(f (Suc i) u * (z-u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) /  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1106  | 
of_nat (fact i)) =  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1107  | 
f (Suc 0) u -  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1108  | 
(f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) /  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1109  | 
of_nat (fact (Suc n)) +  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1110  | 
(\<Sum>i = 0..n.  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1111  | 
(f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) /  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1112  | 
of_nat (fact (Suc i)))"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1113  | 
by (subst setsum_Suc_reindex) simp  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1114  | 
also have "... = f (Suc 0) u -  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1115  | 
(f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) /  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1116  | 
of_nat (fact (Suc n)) +  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1117  | 
(\<Sum>i = 0..n.  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1118  | 
f (Suc (Suc i)) u * ((z-u) ^ Suc i) / of_nat (fact (Suc i)) -  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1119  | 
f (Suc i) u * (z-u) ^ i / of_nat (fact i))"  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
56889 
diff
changeset
 | 
1120  | 
by (simp only: diff_divide_distrib fact_cancel ac_simps)  | 
| 
56238
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1121  | 
also have "... = f (Suc 0) u -  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1122  | 
(f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) /  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1123  | 
of_nat (fact (Suc n)) +  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1124  | 
f (Suc (Suc n)) u * (z-u) ^ Suc n / of_nat (fact (Suc n)) - f (Suc 0) u"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1125  | 
by (subst setsum_Suc_diff) auto  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1126  | 
also have "... = f (Suc n) u * (z-u) ^ n / of_nat (fact n)"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1127  | 
by (simp only: algebra_simps diff_divide_distrib fact_cancel)  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1128  | 
finally have "(\<Sum>i = 0..n. (f (Suc i) u * (z - u) ^ i  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1129  | 
- of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / of_nat (fact i)) =  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1130  | 
f (Suc n) u * (z - u) ^ n / of_nat (fact n)" .  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1131  | 
then have "((\<lambda>u. \<Sum>i = 0..n. f i u * (z - u) ^ i / of_nat (fact i)) has_field_derivative  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1132  | 
f (Suc n) u * (z - u) ^ n / of_nat (fact n)) (at u)"  | 
| 
56381
 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 
hoelzl 
parents: 
56371 
diff
changeset
 | 
1133  | 
apply (intro derivative_eq_intros)+  | 
| 
56238
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1134  | 
apply (force intro: u assms)  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1135  | 
apply (rule refl)+  | 
| 
57514
 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 
haftmann 
parents: 
56889 
diff
changeset
 | 
1136  | 
apply (auto simp: ac_simps)  | 
| 
56238
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1137  | 
done  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1138  | 
}  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1139  | 
then show ?thesis  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1140  | 
apply (cut_tac complex_mvt_line [of w z "\<lambda>u. \<Sum>i = 0..n. f i u * (z-u) ^ i / of_nat (fact i)"  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1141  | 
"\<lambda>u. (f (Suc n) u * (z-u)^n / of_nat (fact n))"])  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1142  | 
apply (auto simp add: intro: open_closed_segment)  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1143  | 
done  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1144  | 
qed  | 
| 
 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 
paulson <lp15@cam.ac.uk> 
parents: 
56223 
diff
changeset
 | 
1145  | 
|
| 56215 | 1146  | 
end  |