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