section ‹Complex Analysis Basics›
theory Complex_Analysis_Basics
imports Cartesian_Euclidean_Space "~~/src/HOL/Library/Nonpos_Ints"
begin
subsection‹General lemmas›
lemma nonneg_Reals_cmod_eq_Re: "z ∈ ℝ⇩≥⇩0 ⟹ norm z = Re z"
by (simp add: complex_nonneg_Reals_iff cmod_eq_Re)
lemma has_derivative_mult_right:
fixes c:: "'a :: real_normed_algebra"
shows "((op * c) has_derivative (op * c)) F"
by (rule has_derivative_mult_right [OF has_derivative_id])
lemma has_derivative_of_real[derivative_intros, simp]:
"(f has_derivative f') F ⟹ ((λx. of_real (f x)) has_derivative (λx. of_real (f' x))) F"
using bounded_linear.has_derivative[OF bounded_linear_of_real] .
lemma has_vector_derivative_real_complex:
"DERIV f (of_real a) :> f' ⟹ ((λx. f (of_real x)) has_vector_derivative f') (at a within s)"
using has_derivative_compose[of of_real of_real a _ f "op * f'"]
by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def)
lemma fact_cancel:
fixes c :: "'a::real_field"
shows "of_nat (Suc n) * c / (fact (Suc n)) = c / (fact n)"
by (simp add: of_nat_mult del: of_nat_Suc times_nat.simps)
lemma bilinear_times:
fixes c::"'a::real_algebra" shows "bilinear (λx y::'a. x*y)"
by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI)
lemma linear_cnj: "linear cnj"
using bounded_linear.linear[OF bounded_linear_cnj] .
lemma tendsto_Re_upper:
assumes "~ (trivial_limit F)"
"(f ⤏ l) F"
"eventually (λx. Re(f x) ≤ b) F"
shows "Re(l) ≤ b"
by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Re)
lemma tendsto_Re_lower:
assumes "~ (trivial_limit F)"
"(f ⤏ l) F"
"eventually (λx. b ≤ Re(f x)) F"
shows "b ≤ Re(l)"
by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Re)
lemma tendsto_Im_upper:
assumes "~ (trivial_limit F)"
"(f ⤏ l) F"
"eventually (λx. Im(f x) ≤ b) F"
shows "Im(l) ≤ b"
by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Im)
lemma tendsto_Im_lower:
assumes "~ (trivial_limit F)"
"(f ⤏ l) F"
"eventually (λx. b ≤ Im(f x)) F"
shows "b ≤ Im(l)"
by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Im)
lemma lambda_zero: "(λh::'a::mult_zero. 0) = op * 0"
by auto
lemma lambda_one: "(λx::'a::monoid_mult. x) = op * 1"
by auto
lemma continuous_mult_left:
fixes c::"'a::real_normed_algebra"
shows "continuous F f ⟹ continuous F (λx. c * f x)"
by (rule continuous_mult [OF continuous_const])
lemma continuous_mult_right:
fixes c::"'a::real_normed_algebra"
shows "continuous F f ⟹ continuous F (λx. f x * c)"
by (rule continuous_mult [OF _ continuous_const])
lemma continuous_on_mult_left:
fixes c::"'a::real_normed_algebra"
shows "continuous_on s f ⟹ continuous_on s (λx. c * f x)"
by (rule continuous_on_mult [OF continuous_on_const])
lemma continuous_on_mult_right:
fixes c::"'a::real_normed_algebra"
shows "continuous_on s f ⟹ continuous_on s (λx. f x * c)"
by (rule continuous_on_mult [OF _ continuous_on_const])
lemma uniformly_continuous_on_cmul_right [continuous_intros]:
fixes f :: "'a::real_normed_vector ⇒ 'b::real_normed_algebra"
shows "uniformly_continuous_on s f ⟹ uniformly_continuous_on s (λx. f x * c)"
using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] .
lemma uniformly_continuous_on_cmul_left[continuous_intros]:
fixes f :: "'a::real_normed_vector ⇒ 'b::real_normed_algebra"
assumes "uniformly_continuous_on s f"
shows "uniformly_continuous_on s (λx. c * f x)"
by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right)
lemma continuous_within_norm_id [continuous_intros]: "continuous (at x within S) norm"
by (rule continuous_norm [OF continuous_ident])
lemma continuous_on_norm_id [continuous_intros]: "continuous_on S norm"
by (intro continuous_on_id continuous_on_norm)
subsection‹DERIV stuff›
lemma DERIV_zero_connected_constant:
fixes f :: "'a::{real_normed_field,euclidean_space} ⇒ 'a"
assumes "connected s"
and "open s"
and "finite k"
and "continuous_on s f"
and "∀x∈(s - k). DERIV f x :> 0"
obtains c where "⋀x. x ∈ s ⟹ f(x) = c"
using has_derivative_zero_connected_constant [OF assms(1-4)] assms
by (metis DERIV_const has_derivative_const Diff_iff at_within_open frechet_derivative_at has_field_derivative_def)
lemma DERIV_zero_constant:
fixes f :: "'a::{real_normed_field, real_inner} ⇒ 'a"
shows "⟦convex s;
⋀x. x∈s ⟹ (f has_field_derivative 0) (at x within s)⟧
⟹ ∃c. ∀x ∈ s. f(x) = c"
by (auto simp: has_field_derivative_def lambda_zero intro: has_derivative_zero_constant)
lemma DERIV_zero_unique:
fixes f :: "'a::{real_normed_field, real_inner} ⇒ 'a"
assumes "convex s"
and d0: "⋀x. x∈s ⟹ (f has_field_derivative 0) (at x within s)"
and "a ∈ s"
and "x ∈ s"
shows "f x = f a"
by (rule has_derivative_zero_unique [OF assms(1) _ assms(4,3)])
(metis d0 has_field_derivative_imp_has_derivative lambda_zero)
lemma DERIV_zero_connected_unique:
fixes f :: "'a::{real_normed_field, real_inner} ⇒ 'a"
assumes "connected s"
and "open s"
and d0: "⋀x. x∈s ⟹ DERIV f x :> 0"
and "a ∈ s"
and "x ∈ s"
shows "f x = f a"
by (rule has_derivative_zero_unique_connected [OF assms(2,1) _ assms(5,4)])
(metis has_field_derivative_def lambda_zero d0)
lemma DERIV_transform_within:
assumes "(f has_field_derivative f') (at a within s)"
and "0 < d" "a ∈ s"
and "⋀x. x∈s ⟹ dist x a < d ⟹ f x = g x"
shows "(g has_field_derivative f') (at a within s)"
using assms unfolding has_field_derivative_def
by (blast intro: has_derivative_transform_within)
lemma DERIV_transform_within_open:
assumes "DERIV f a :> f'"
and "open s" "a ∈ s"
and "⋀x. x∈s ⟹ f x = g x"
shows "DERIV g a :> f'"
using assms unfolding has_field_derivative_def
by (metis has_derivative_transform_within_open)
lemma DERIV_transform_at:
assumes "DERIV f a :> f'"
and "0 < d"
and "⋀x. dist x a < d ⟹ f x = g x"
shows "DERIV g a :> f'"
by (blast intro: assms DERIV_transform_within)
lemma DERIV_zero_UNIV_unique:
fixes f :: "'a::{real_normed_field, real_inner} ⇒ 'a"
shows "(⋀x. DERIV f x :> 0) ⟹ f x = f a"
by (metis DERIV_zero_unique UNIV_I assms convex_UNIV)
subsection ‹Some limit theorems about real part of real series etc.›
lemma sums_vec_nth :
assumes "f sums a"
shows "(λx. f x $ i) sums a $ i"
using assms unfolding sums_def
by (auto dest: tendsto_vec_nth [where i=i])
lemma summable_vec_nth :
assumes "summable f"
shows "summable (λx. f x $ i)"
using assms unfolding summable_def
by (blast intro: sums_vec_nth)
subsection ‹Complex number lemmas›
lemma
shows open_halfspace_Re_lt: "open {z. Re(z) < b}"
and open_halfspace_Re_gt: "open {z. Re(z) > b}"
and closed_halfspace_Re_ge: "closed {z. Re(z) ≥ b}"
and closed_halfspace_Re_le: "closed {z. Re(z) ≤ b}"
and closed_halfspace_Re_eq: "closed {z. Re(z) = b}"
and open_halfspace_Im_lt: "open {z. Im(z) < b}"
and open_halfspace_Im_gt: "open {z. Im(z) > b}"
and closed_halfspace_Im_ge: "closed {z. Im(z) ≥ b}"
and closed_halfspace_Im_le: "closed {z. Im(z) ≤ b}"
and closed_halfspace_Im_eq: "closed {z. Im(z) = b}"
by (intro open_Collect_less closed_Collect_le closed_Collect_eq isCont_Re
isCont_Im continuous_ident continuous_const)+
lemma closed_complex_Reals: "closed (ℝ :: complex set)"
proof -
have "(ℝ :: complex set) = {z. Im z = 0}"
by (auto simp: complex_is_Real_iff)
then show ?thesis
by (metis closed_halfspace_Im_eq)
qed
lemma closed_Real_halfspace_Re_le: "closed (ℝ ∩ {w. Re w ≤ x})"
by (simp add: closed_Int closed_complex_Reals closed_halfspace_Re_le)
corollary closed_nonpos_Reals_complex [simp]: "closed (ℝ⇩≤⇩0 :: complex set)"
proof -
have "ℝ⇩≤⇩0 = ℝ ∩ {z. Re(z) ≤ 0}"
using complex_nonpos_Reals_iff complex_is_Real_iff by auto
then show ?thesis
by (metis closed_Real_halfspace_Re_le)
qed
lemma closed_Real_halfspace_Re_ge: "closed (ℝ ∩ {w. x ≤ Re(w)})"
using closed_halfspace_Re_ge
by (simp add: closed_Int closed_complex_Reals)
corollary closed_nonneg_Reals_complex [simp]: "closed (ℝ⇩≥⇩0 :: complex set)"
proof -
have "ℝ⇩≥⇩0 = ℝ ∩ {z. Re(z) ≥ 0}"
using complex_nonneg_Reals_iff complex_is_Real_iff by auto
then show ?thesis
by (metis closed_Real_halfspace_Re_ge)
qed
lemma closed_real_abs_le: "closed {w ∈ ℝ. ¦Re w¦ ≤ r}"
proof -
have "{w ∈ ℝ. ¦Re w¦ ≤ r} = (ℝ ∩ {w. Re w ≤ r}) ∩ (ℝ ∩ {w. Re w ≥ -r})"
by auto
then show "closed {w ∈ ℝ. ¦Re w¦ ≤ r}"
by (simp add: closed_Int closed_Real_halfspace_Re_ge closed_Real_halfspace_Re_le)
qed
lemma real_lim:
fixes l::complex
assumes "(f ⤏ l) F" and "~(trivial_limit F)" and "eventually P F" and "⋀a. P a ⟹ f a ∈ ℝ"
shows "l ∈ ℝ"
proof (rule Lim_in_closed_set[OF closed_complex_Reals _ assms(2,1)])
show "eventually (λx. f x ∈ ℝ) F"
using assms(3, 4) by (auto intro: eventually_mono)
qed
lemma real_lim_sequentially:
fixes l::complex
shows "(f ⤏ l) sequentially ⟹ (∃N. ∀n≥N. f n ∈ ℝ) ⟹ l ∈ ℝ"
by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially)
lemma real_series:
fixes l::complex
shows "f sums l ⟹ (⋀n. f n ∈ ℝ) ⟹ l ∈ ℝ"
unfolding sums_def
by (metis real_lim_sequentially setsum_in_Reals)
lemma Lim_null_comparison_Re:
assumes "eventually (λx. norm(f x) ≤ Re(g x)) F" "(g ⤏ 0) F" shows "(f ⤏ 0) F"
by (rule Lim_null_comparison[OF assms(1)] tendsto_eq_intros assms(2))+ simp
subsection‹Holomorphic functions›
definition field_differentiable :: "['a ⇒ 'a::real_normed_field, 'a filter] ⇒ bool"
(infixr "(field'_differentiable)" 50)
where "f field_differentiable F ≡ ∃f'. (f has_field_derivative f') F"
lemma field_differentiable_derivI:
"f field_differentiable (at x) ⟹ (f has_field_derivative deriv f x) (at x)"
by (simp add: field_differentiable_def DERIV_deriv_iff_has_field_derivative)
lemma field_differentiable_imp_continuous_at:
"f field_differentiable (at x within s) ⟹ continuous (at x within s) f"
by (metis DERIV_continuous field_differentiable_def)
lemma field_differentiable_within_subset:
"⟦f field_differentiable (at x within s); t ⊆ s⟧
⟹ f field_differentiable (at x within t)"
by (metis DERIV_subset field_differentiable_def)
lemma field_differentiable_at_within:
"⟦f field_differentiable (at x)⟧
⟹ f field_differentiable (at x within s)"
unfolding field_differentiable_def
by (metis DERIV_subset top_greatest)
lemma field_differentiable_linear [simp,derivative_intros]: "(op * c) field_differentiable F"
proof -
show ?thesis
unfolding field_differentiable_def has_field_derivative_def mult_commute_abs
by (force intro: has_derivative_mult_right)
qed
lemma field_differentiable_const [simp,derivative_intros]: "(λz. c) field_differentiable F"
unfolding field_differentiable_def has_field_derivative_def
by (rule exI [where x=0])
(metis has_derivative_const lambda_zero)
lemma field_differentiable_ident [simp,derivative_intros]: "(λz. z) field_differentiable F"
unfolding field_differentiable_def has_field_derivative_def
by (rule exI [where x=1])
(simp add: lambda_one [symmetric])
lemma field_differentiable_id [simp,derivative_intros]: "id field_differentiable F"
unfolding id_def by (rule field_differentiable_ident)
lemma field_differentiable_minus [derivative_intros]:
"f field_differentiable F ⟹ (λz. - (f z)) field_differentiable F"
using assms unfolding field_differentiable_def
by (metis field_differentiable_minus)
lemma field_differentiable_add [derivative_intros]:
assumes "f field_differentiable F" "g field_differentiable F"
shows "(λz. f z + g z) field_differentiable F"
using assms unfolding field_differentiable_def
by (metis field_differentiable_add)
lemma field_differentiable_add_const [simp,derivative_intros]:
"op + c field_differentiable F"
by (simp add: field_differentiable_add)
lemma field_differentiable_setsum [derivative_intros]:
"(⋀i. i ∈ I ⟹ (f i) field_differentiable F) ⟹ (λz. ∑i∈I. f i z) field_differentiable F"
by (induct I rule: infinite_finite_induct)
(auto intro: field_differentiable_add field_differentiable_const)
lemma field_differentiable_diff [derivative_intros]:
assumes "f field_differentiable F" "g field_differentiable F"
shows "(λz. f z - g z) field_differentiable F"
using assms unfolding field_differentiable_def
by (metis field_differentiable_diff)
lemma field_differentiable_inverse [derivative_intros]:
assumes "f field_differentiable (at a within s)" "f a ≠ 0"
shows "(λz. inverse (f z)) field_differentiable (at a within s)"
using assms unfolding field_differentiable_def
by (metis DERIV_inverse_fun)
lemma field_differentiable_mult [derivative_intros]:
assumes "f field_differentiable (at a within s)"
"g field_differentiable (at a within s)"
shows "(λz. f z * g z) field_differentiable (at a within s)"
using assms unfolding field_differentiable_def
by (metis DERIV_mult [of f _ a s g])
lemma field_differentiable_divide [derivative_intros]:
assumes "f field_differentiable (at a within s)"
"g field_differentiable (at a within s)"
"g a ≠ 0"
shows "(λz. f z / g z) field_differentiable (at a within s)"
using assms unfolding field_differentiable_def
by (metis DERIV_divide [of f _ a s g])
lemma field_differentiable_power [derivative_intros]:
assumes "f field_differentiable (at a within s)"
shows "(λz. f z ^ n) field_differentiable (at a within s)"
using assms unfolding field_differentiable_def
by (metis DERIV_power)
lemma field_differentiable_transform_within:
"0 < d ⟹
x ∈ s ⟹
(⋀x'. x' ∈ s ⟹ dist x' x < d ⟹ f x' = g x') ⟹
f field_differentiable (at x within s)
⟹ g field_differentiable (at x within s)"
unfolding field_differentiable_def has_field_derivative_def
by (blast intro: has_derivative_transform_within)
lemma field_differentiable_compose_within:
assumes "f field_differentiable (at a within s)"
"g field_differentiable (at (f a) within f`s)"
shows "(g o f) field_differentiable (at a within s)"
using assms unfolding field_differentiable_def
by (metis DERIV_image_chain)
lemma field_differentiable_compose:
"f field_differentiable at z ⟹ g field_differentiable at (f z)
⟹ (g o f) field_differentiable at z"
by (metis field_differentiable_at_within field_differentiable_compose_within)
lemma field_differentiable_within_open:
"⟦a ∈ s; open s⟧ ⟹ f field_differentiable at a within s ⟷
f field_differentiable at a"
unfolding field_differentiable_def
by (metis at_within_open)
subsection‹Caratheodory characterization›
lemma field_differentiable_caratheodory_at:
"f field_differentiable (at z) ⟷
(∃g. (∀w. f(w) - f(z) = g(w) * (w - z)) ∧ continuous (at z) g)"
using CARAT_DERIV [of f]
by (simp add: field_differentiable_def has_field_derivative_def)
lemma field_differentiable_caratheodory_within:
"f field_differentiable (at z within s) ⟷
(∃g. (∀w. f(w) - f(z) = g(w) * (w - z)) ∧ continuous (at z within s) g)"
using DERIV_caratheodory_within [of f]
by (simp add: field_differentiable_def has_field_derivative_def)
subsection‹Holomorphic›
definition holomorphic_on :: "[complex ⇒ complex, complex set] ⇒ bool"
(infixl "(holomorphic'_on)" 50)
where "f holomorphic_on s ≡ ∀x∈s. f field_differentiable (at x within s)"
named_theorems holomorphic_intros "structural introduction rules for holomorphic_on"
lemma holomorphic_onI [intro?]: "(⋀x. x ∈ s ⟹ f field_differentiable (at x within s)) ⟹ f holomorphic_on s"
by (simp add: holomorphic_on_def)
lemma holomorphic_onD [dest?]: "⟦f holomorphic_on s; x ∈ s⟧ ⟹ f field_differentiable (at x within s)"
by (simp add: holomorphic_on_def)
lemma holomorphic_on_imp_differentiable_at:
"⟦f holomorphic_on s; open s; x ∈ s⟧ ⟹ f field_differentiable (at x)"
using at_within_open holomorphic_on_def by fastforce
lemma holomorphic_on_empty [holomorphic_intros]: "f holomorphic_on {}"
by (simp add: holomorphic_on_def)
lemma holomorphic_on_open:
"open s ⟹ f holomorphic_on s ⟷ (∀x ∈ s. ∃f'. DERIV f x :> f')"
by (auto simp: holomorphic_on_def field_differentiable_def has_field_derivative_def at_within_open [of _ s])
lemma holomorphic_on_imp_continuous_on:
"f holomorphic_on s ⟹ continuous_on s f"
by (metis field_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def)
lemma holomorphic_on_subset [elim]:
"f holomorphic_on s ⟹ t ⊆ s ⟹ f holomorphic_on t"
unfolding holomorphic_on_def
by (metis field_differentiable_within_subset subsetD)
lemma holomorphic_transform: "⟦f holomorphic_on s; ⋀x. x ∈ s ⟹ f x = g x⟧ ⟹ g holomorphic_on s"
by (metis field_differentiable_transform_within linordered_field_no_ub holomorphic_on_def)
lemma holomorphic_cong: "s = t ==> (⋀x. x ∈ s ⟹ f x = g x) ⟹ f holomorphic_on s ⟷ g holomorphic_on t"
by (metis holomorphic_transform)
lemma holomorphic_on_linear [simp, holomorphic_intros]: "(op * c) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_linear)
lemma holomorphic_on_const [simp, holomorphic_intros]: "(λz. c) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_const)
lemma holomorphic_on_ident [simp, holomorphic_intros]: "(λx. x) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_ident)
lemma holomorphic_on_id [simp, holomorphic_intros]: "id holomorphic_on s"
unfolding id_def by (rule holomorphic_on_ident)
lemma holomorphic_on_compose:
"f holomorphic_on s ⟹ g holomorphic_on (f ` s) ⟹ (g o f) holomorphic_on s"
using field_differentiable_compose_within[of f _ s g]
by (auto simp: holomorphic_on_def)
lemma holomorphic_on_compose_gen:
"f holomorphic_on s ⟹ g holomorphic_on t ⟹ f ` s ⊆ t ⟹ (g o f) holomorphic_on s"
by (metis holomorphic_on_compose holomorphic_on_subset)
lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s ⟹ (λz. -(f z)) holomorphic_on s"
by (metis field_differentiable_minus holomorphic_on_def)
lemma holomorphic_on_add [holomorphic_intros]:
"⟦f holomorphic_on s; g holomorphic_on s⟧ ⟹ (λz. f z + g z) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_add)
lemma holomorphic_on_diff [holomorphic_intros]:
"⟦f holomorphic_on s; g holomorphic_on s⟧ ⟹ (λz. f z - g z) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_diff)
lemma holomorphic_on_mult [holomorphic_intros]:
"⟦f holomorphic_on s; g holomorphic_on s⟧ ⟹ (λz. f z * g z) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_mult)
lemma holomorphic_on_inverse [holomorphic_intros]:
"⟦f holomorphic_on s; ⋀z. z ∈ s ⟹ f z ≠ 0⟧ ⟹ (λz. inverse (f z)) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_inverse)
lemma holomorphic_on_divide [holomorphic_intros]:
"⟦f holomorphic_on s; g holomorphic_on s; ⋀z. z ∈ s ⟹ g z ≠ 0⟧ ⟹ (λz. f z / g z) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_divide)
lemma holomorphic_on_power [holomorphic_intros]:
"f holomorphic_on s ⟹ (λz. (f z)^n) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_power)
lemma holomorphic_on_setsum [holomorphic_intros]:
"(⋀i. i ∈ I ⟹ (f i) holomorphic_on s) ⟹ (λx. setsum (λi. f i x) I) holomorphic_on s"
unfolding holomorphic_on_def by (metis field_differentiable_setsum)
lemma DERIV_deriv_iff_field_differentiable:
"DERIV f x :> deriv f x ⟷ f field_differentiable at x"
unfolding field_differentiable_def by (metis DERIV_imp_deriv)
lemma holomorphic_derivI:
"⟦f holomorphic_on S; open S; x ∈ S⟧
⟹ (f has_field_derivative deriv f x) (at x within T)"
by (metis DERIV_deriv_iff_field_differentiable at_within_open holomorphic_on_def has_field_derivative_at_within)
lemma complex_derivative_chain:
"f field_differentiable at x ⟹ g field_differentiable at (f x)
⟹ deriv (g o f) x = deriv g (f x) * deriv f x"
by (metis DERIV_deriv_iff_field_differentiable DERIV_chain DERIV_imp_deriv)
lemma deriv_linear [simp]: "deriv (λw. c * w) = (λz. c)"
by (metis DERIV_imp_deriv DERIV_cmult_Id)
lemma deriv_ident [simp]: "deriv (λw. w) = (λz. 1)"
by (metis DERIV_imp_deriv DERIV_ident)
lemma deriv_id [simp]: "deriv id = (λz. 1)"
by (simp add: id_def)
lemma deriv_const [simp]: "deriv (λw. c) = (λz. 0)"
by (metis DERIV_imp_deriv DERIV_const)
lemma deriv_add [simp]:
"⟦f field_differentiable at z; g field_differentiable at z⟧
⟹ deriv (λw. f w + g w) z = deriv f z + deriv g z"
unfolding DERIV_deriv_iff_field_differentiable[symmetric]
by (auto intro!: DERIV_imp_deriv derivative_intros)
lemma deriv_diff [simp]:
"⟦f field_differentiable at z; g field_differentiable at z⟧
⟹ deriv (λw. f w - g w) z = deriv f z - deriv g z"
unfolding DERIV_deriv_iff_field_differentiable[symmetric]
by (auto intro!: DERIV_imp_deriv derivative_intros)
lemma deriv_mult [simp]:
"⟦f field_differentiable at z; g field_differentiable at z⟧
⟹ deriv (λw. f w * g w) z = f z * deriv g z + deriv f z * g z"
unfolding DERIV_deriv_iff_field_differentiable[symmetric]
by (auto intro!: DERIV_imp_deriv derivative_eq_intros)
lemma deriv_cmult [simp]:
"f field_differentiable at z ⟹ deriv (λw. c * f w) z = c * deriv f z"
unfolding DERIV_deriv_iff_field_differentiable[symmetric]
by (auto intro!: DERIV_imp_deriv derivative_eq_intros)
lemma deriv_cmult_right [simp]:
"f field_differentiable at z ⟹ deriv (λw. f w * c) z = deriv f z * c"
unfolding DERIV_deriv_iff_field_differentiable[symmetric]
by (auto intro!: DERIV_imp_deriv derivative_eq_intros)
lemma deriv_cdivide_right [simp]:
"f field_differentiable at z ⟹ deriv (λw. f w / c) z = deriv f z / c"
unfolding Fields.field_class.field_divide_inverse
by (blast intro: deriv_cmult_right)
lemma complex_derivative_transform_within_open:
"⟦f holomorphic_on s; g holomorphic_on s; open s; z ∈ s; ⋀w. w ∈ s ⟹ f w = g w⟧
⟹ deriv f z = deriv g z"
unfolding holomorphic_on_def
by (rule DERIV_imp_deriv)
(metis DERIV_deriv_iff_field_differentiable DERIV_transform_within_open at_within_open)
lemma deriv_compose_linear:
"f field_differentiable at (c * z) ⟹ deriv (λw. f (c * w)) z = c * deriv f (c * z)"
apply (rule DERIV_imp_deriv)
apply (simp add: DERIV_deriv_iff_field_differentiable [symmetric])
apply (drule DERIV_chain' [of "times c" c z UNIV f "deriv f (c * z)", OF DERIV_cmult_Id])
apply (simp add: algebra_simps)
done
lemma nonzero_deriv_nonconstant:
assumes df: "DERIV f ξ :> df" and S: "open S" "ξ ∈ S" and "df ≠ 0"
shows "¬ f constant_on S"
unfolding constant_on_def
by (metis ‹df ≠ 0› DERIV_transform_within_open [OF df S] DERIV_const DERIV_unique)
lemma holomorphic_nonconstant:
assumes holf: "f holomorphic_on S" and "open S" "ξ ∈ S" "deriv f ξ ≠ 0"
shows "¬ f constant_on S"
apply (rule nonzero_deriv_nonconstant [of f "deriv f ξ" ξ S])
using assms
apply (auto simp: holomorphic_derivI)
done
subsection‹Analyticity on a set›
definition analytic_on (infixl "(analytic'_on)" 50)
where
"f analytic_on s ≡ ∀x ∈ s. ∃e. 0 < e ∧ f holomorphic_on (ball x e)"
lemma analytic_imp_holomorphic: "f analytic_on s ⟹ f holomorphic_on s"
by (simp add: at_within_open [OF _ open_ball] analytic_on_def holomorphic_on_def)
(metis centre_in_ball field_differentiable_at_within)
lemma analytic_on_open: "open s ⟹ f analytic_on s ⟷ f holomorphic_on s"
apply (auto simp: analytic_imp_holomorphic)
apply (auto simp: analytic_on_def holomorphic_on_def)
by (metis holomorphic_on_def holomorphic_on_subset open_contains_ball)
lemma analytic_on_imp_differentiable_at:
"f analytic_on s ⟹ x ∈ s ⟹ f field_differentiable (at x)"
apply (auto simp: analytic_on_def holomorphic_on_def)
by (metis Topology_Euclidean_Space.open_ball centre_in_ball field_differentiable_within_open)
lemma analytic_on_subset: "f analytic_on s ⟹ t ⊆ s ⟹ f analytic_on t"
by (auto simp: analytic_on_def)
lemma analytic_on_Un: "f analytic_on (s ∪ t) ⟷ f analytic_on s ∧ f analytic_on t"
by (auto simp: analytic_on_def)
lemma analytic_on_Union: "f analytic_on (⋃s) ⟷ (∀t ∈ s. f analytic_on t)"
by (auto simp: analytic_on_def)
lemma analytic_on_UN: "f analytic_on (⋃i∈I. s i) ⟷ (∀i∈I. f analytic_on (s i))"
by (auto simp: analytic_on_def)
lemma analytic_on_holomorphic:
"f analytic_on s ⟷ (∃t. open t ∧ s ⊆ t ∧ f holomorphic_on t)"
(is "?lhs = ?rhs")
proof -
have "?lhs ⟷ (∃t. open t ∧ s ⊆ t ∧ f analytic_on t)"
proof safe
assume "f analytic_on s"
then show "∃t. open t ∧ s ⊆ t ∧ f analytic_on t"
apply (simp add: analytic_on_def)
apply (rule exI [where x="⋃{u. open u ∧ f analytic_on u}"], auto)
apply (metis Topology_Euclidean_Space.open_ball analytic_on_open centre_in_ball)
by (metis analytic_on_def)
next
fix t
assume "open t" "s ⊆ t" "f analytic_on t"
then show "f analytic_on s"
by (metis analytic_on_subset)
qed
also have "... ⟷ ?rhs"
by (auto simp: analytic_on_open)
finally show ?thesis .
qed
lemma analytic_on_linear: "(op * c) analytic_on s"
by (auto simp add: analytic_on_holomorphic holomorphic_on_linear)
lemma analytic_on_const: "(λz. c) analytic_on s"
by (metis analytic_on_def holomorphic_on_const zero_less_one)
lemma analytic_on_ident: "(λx. x) analytic_on s"
by (simp add: analytic_on_def holomorphic_on_ident gt_ex)
lemma analytic_on_id: "id analytic_on s"
unfolding id_def by (rule analytic_on_ident)
lemma analytic_on_compose:
assumes f: "f analytic_on s"
and g: "g analytic_on (f ` s)"
shows "(g o f) analytic_on s"
unfolding analytic_on_def
proof (intro ballI)
fix x
assume x: "x ∈ s"
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f
by (metis analytic_on_def)
obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g
by (metis analytic_on_def g image_eqI x)
have "isCont f x"
by (metis analytic_on_imp_differentiable_at field_differentiable_imp_continuous_at f x)
with e' obtain d where d: "0 < d" and fd: "f ` ball x d ⊆ ball (f x) e'"
by (auto simp: continuous_at_ball)
have "g ∘ f holomorphic_on ball x (min d e)"
apply (rule holomorphic_on_compose)
apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball)
then show "∃e>0. g ∘ f holomorphic_on ball x e"
by (metis d e min_less_iff_conj)
qed
lemma analytic_on_compose_gen:
"f analytic_on s ⟹ g analytic_on t ⟹ (⋀z. z ∈ s ⟹ f z ∈ t)
⟹ g o f analytic_on s"
by (metis analytic_on_compose analytic_on_subset image_subset_iff)
lemma analytic_on_neg:
"f analytic_on s ⟹ (λz. -(f z)) analytic_on s"
by (metis analytic_on_holomorphic holomorphic_on_minus)
lemma analytic_on_add:
assumes f: "f analytic_on s"
and g: "g analytic_on s"
shows "(λz. f z + g z) analytic_on s"
unfolding analytic_on_def
proof (intro ballI)
fix z
assume z: "z ∈ s"
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
by (metis analytic_on_def)
obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
by (metis analytic_on_def g z)
have "(λz. f z + g z) holomorphic_on ball z (min e e')"
apply (rule holomorphic_on_add)
apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
then show "∃e>0. (λz. f z + g z) holomorphic_on ball z e"
by (metis e e' min_less_iff_conj)
qed
lemma analytic_on_diff:
assumes f: "f analytic_on s"
and g: "g analytic_on s"
shows "(λz. f z - g z) analytic_on s"
unfolding analytic_on_def
proof (intro ballI)
fix z
assume z: "z ∈ s"
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
by (metis analytic_on_def)
obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
by (metis analytic_on_def g z)
have "(λz. f z - g z) holomorphic_on ball z (min e e')"
apply (rule holomorphic_on_diff)
apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
then show "∃e>0. (λz. f z - g z) holomorphic_on ball z e"
by (metis e e' min_less_iff_conj)
qed
lemma analytic_on_mult:
assumes f: "f analytic_on s"
and g: "g analytic_on s"
shows "(λz. f z * g z) analytic_on s"
unfolding analytic_on_def
proof (intro ballI)
fix z
assume z: "z ∈ s"
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
by (metis analytic_on_def)
obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g
by (metis analytic_on_def g z)
have "(λz. f z * g z) holomorphic_on ball z (min e e')"
apply (rule holomorphic_on_mult)
apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball)
then show "∃e>0. (λz. f z * g z) holomorphic_on ball z e"
by (metis e e' min_less_iff_conj)
qed
lemma analytic_on_inverse:
assumes f: "f analytic_on s"
and nz: "(⋀z. z ∈ s ⟹ f z ≠ 0)"
shows "(λz. inverse (f z)) analytic_on s"
unfolding analytic_on_def
proof (intro ballI)
fix z
assume z: "z ∈ s"
then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f
by (metis analytic_on_def)
have "continuous_on (ball z e) f"
by (metis fh holomorphic_on_imp_continuous_on)
then obtain e' where e': "0 < e'" and nz': "⋀y. dist z y < e' ⟹ f y ≠ 0"
by (metis Topology_Euclidean_Space.open_ball centre_in_ball continuous_on_open_avoid e z nz)
have "(λz. inverse (f z)) holomorphic_on ball z (min e e')"
apply (rule holomorphic_on_inverse)
apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball)
by (metis nz' mem_ball min_less_iff_conj)
then show "∃e>0. (λz. inverse (f z)) holomorphic_on ball z e"
by (metis e e' min_less_iff_conj)
qed
lemma analytic_on_divide:
assumes f: "f analytic_on s"
and g: "g analytic_on s"
and nz: "(⋀z. z ∈ s ⟹ g z ≠ 0)"
shows "(λz. f z / g z) analytic_on s"
unfolding divide_inverse
by (metis analytic_on_inverse analytic_on_mult f g nz)
lemma analytic_on_power:
"f analytic_on s ⟹ (λz. (f z) ^ n) analytic_on s"
by (induct n) (auto simp: analytic_on_const analytic_on_mult)
lemma analytic_on_setsum:
"(⋀i. i ∈ I ⟹ (f i) analytic_on s) ⟹ (λx. setsum (λi. f i x) I) analytic_on s"
by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add)
lemma deriv_left_inverse:
assumes "f holomorphic_on S" and "g holomorphic_on T"
and "open S" and "open T"
and "f ` S ⊆ T"
and [simp]: "⋀z. z ∈ S ⟹ g (f z) = z"
and "w ∈ S"
shows "deriv f w * deriv g (f w) = 1"
proof -
have "deriv f w * deriv g (f w) = deriv g (f w) * deriv f w"
by (simp add: algebra_simps)
also have "... = deriv (g o f) w"
using assms
by (metis analytic_on_imp_differentiable_at analytic_on_open complex_derivative_chain image_subset_iff)
also have "... = deriv id w"
apply (rule complex_derivative_transform_within_open [where s=S])
apply (rule assms holomorphic_on_compose_gen holomorphic_intros)+
apply simp
done
also have "... = 1"
by simp
finally show ?thesis .
qed
subsection‹analyticity at a point›
lemma analytic_at_ball:
"f analytic_on {z} ⟷ (∃e. 0<e ∧ f holomorphic_on ball z e)"
by (metis analytic_on_def singleton_iff)
lemma analytic_at:
"f analytic_on {z} ⟷ (∃s. open s ∧ z ∈ s ∧ f holomorphic_on s)"
by (metis analytic_on_holomorphic empty_subsetI insert_subset)
lemma analytic_on_analytic_at:
"f analytic_on s ⟷ (∀z ∈ s. f analytic_on {z})"
by (metis analytic_at_ball analytic_on_def)
lemma analytic_at_two:
"f analytic_on {z} ∧ g analytic_on {z} ⟷
(∃s. open s ∧ z ∈ s ∧ f holomorphic_on s ∧ g holomorphic_on s)"
(is "?lhs = ?rhs")
proof
assume ?lhs
then obtain s t
where st: "open s" "z ∈ s" "f holomorphic_on s"
"open t" "z ∈ t" "g holomorphic_on t"
by (auto simp: analytic_at)
show ?rhs
apply (rule_tac x="s ∩ t" in exI)
using st
apply (auto simp: Diff_subset holomorphic_on_subset)
done
next
assume ?rhs
then show ?lhs
by (force simp add: analytic_at)
qed
subsection‹Combining theorems for derivative with ``analytic at'' hypotheses›
lemma
assumes "f analytic_on {z}" "g analytic_on {z}"
shows complex_derivative_add_at: "deriv (λw. f w + g w) z = deriv f z + deriv g z"
and complex_derivative_diff_at: "deriv (λw. f w - g w) z = deriv f z - deriv g z"
and complex_derivative_mult_at: "deriv (λw. f w * g w) z =
f z * deriv g z + deriv f z * g z"
proof -
obtain s where s: "open s" "z ∈ s" "f holomorphic_on s" "g holomorphic_on s"
using assms by (metis analytic_at_two)
show "deriv (λw. f w + g w) z = deriv f z + deriv g z"
apply (rule DERIV_imp_deriv [OF DERIV_add])
using s
apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable)
done
show "deriv (λw. f w - g w) z = deriv f z - deriv g z"
apply (rule DERIV_imp_deriv [OF DERIV_diff])
using s
apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable)
done
show "deriv (λw. f w * g w) z = f z * deriv g z + deriv f z * g z"
apply (rule DERIV_imp_deriv [OF DERIV_mult'])
using s
apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable)
done
qed
lemma deriv_cmult_at:
"f analytic_on {z} ⟹ deriv (λw. c * f w) z = c * deriv f z"
by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const)
lemma deriv_cmult_right_at:
"f analytic_on {z} ⟹ deriv (λw. f w * c) z = deriv f z * c"
by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const)
subsection‹Complex differentiation of sequences and series›
lemma has_complex_derivative_sequence:
fixes s :: "complex set"
assumes cvs: "convex s"
and df: "⋀n x. x ∈ s ⟹ (f n has_field_derivative f' n x) (at x within s)"
and conv: "⋀e. 0 < e ⟹ ∃N. ∀n x. n ≥ N ⟶ x ∈ s ⟶ norm (f' n x - g' x) ≤ e"
and "∃x l. x ∈ s ∧ ((λn. f n x) ⤏ l) sequentially"
shows "∃g. ∀x ∈ s. ((λn. f n x) ⤏ g x) sequentially ∧
(g has_field_derivative (g' x)) (at x within s)"
proof -
from assms obtain x l where x: "x ∈ s" and tf: "((λn. f n x) ⤏ l) sequentially"
by blast
{ fix e::real assume e: "e > 0"
then obtain N where N: "∀n≥N. ∀x. x ∈ s ⟶ cmod (f' n x - g' x) ≤ e"
by (metis conv)
have "∃N. ∀n≥N. ∀x∈s. ∀h. cmod (f' n x * h - g' x * h) ≤ e * cmod h"
proof (rule exI [of _ N], clarify)
fix n y h
assume "N ≤ n" "y ∈ s"
then have "cmod (f' n y - g' y) ≤ e"
by (metis N)
then have "cmod h * cmod (f' n y - g' y) ≤ cmod h * e"
by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)
then show "cmod (f' n y * h - g' y * h) ≤ e * cmod h"
by (simp add: norm_mult [symmetric] field_simps)
qed
} note ** = this
show ?thesis
unfolding has_field_derivative_def
proof (rule has_derivative_sequence [OF cvs _ _ x])
show "∀n. ∀x∈s. (f n has_derivative (op * (f' n x))) (at x within s)"
by (metis has_field_derivative_def df)
next show "(λn. f n x) ⇢ l"
by (rule tf)
next show "∀e>0. ∃N. ∀n≥N. ∀x∈s. ∀h. cmod (f' n x * h - g' x * h) ≤ e * cmod h"
by (blast intro: **)
qed
qed
lemma has_complex_derivative_series:
fixes s :: "complex set"
assumes cvs: "convex s"
and df: "⋀n x. x ∈ s ⟹ (f n has_field_derivative f' n x) (at x within s)"
and conv: "⋀e. 0 < e ⟹ ∃N. ∀n x. n ≥ N ⟶ x ∈ s
⟶ cmod ((∑i<n. f' i x) - g' x) ≤ e"
and "∃x l. x ∈ s ∧ ((λn. f n x) sums l)"
shows "∃g. ∀x ∈ s. ((λn. f n x) sums g x) ∧ ((g has_field_derivative g' x) (at x within s))"
proof -
from assms obtain x l where x: "x ∈ s" and sf: "((λn. f n x) sums l)"
by blast
{ fix e::real assume e: "e > 0"
then obtain N where N: "∀n x. n ≥ N ⟶ x ∈ s
⟶ cmod ((∑i<n. f' i x) - g' x) ≤ e"
by (metis conv)
have "∃N. ∀n≥N. ∀x∈s. ∀h. cmod ((∑i<n. h * f' i x) - g' x * h) ≤ e * cmod h"
proof (rule exI [of _ N], clarify)
fix n y h
assume "N ≤ n" "y ∈ s"
then have "cmod ((∑i<n. f' i y) - g' y) ≤ e"
by (metis N)
then have "cmod h * cmod ((∑i<n. f' i y) - g' y) ≤ cmod h * e"
by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2)
then show "cmod ((∑i<n. h * f' i y) - g' y * h) ≤ e * cmod h"
by (simp add: norm_mult [symmetric] field_simps setsum_right_distrib)
qed
} note ** = this
show ?thesis
unfolding has_field_derivative_def
proof (rule has_derivative_series [OF cvs _ _ x])
fix n x
assume "x ∈ s"
then show "((f n) has_derivative (λz. z * f' n x)) (at x within s)"
by (metis df has_field_derivative_def mult_commute_abs)
next show " ((λn. f n x) sums l)"
by (rule sf)
next show "∀e>0. ∃N. ∀n≥N. ∀x∈s. ∀h. cmod ((∑i<n. h * f' i x) - g' x * h) ≤ e * cmod h"
by (blast intro: **)
qed
qed
lemma field_differentiable_series:
fixes f :: "nat ⇒ complex ⇒ complex"
assumes "convex s" "open s"
assumes "⋀n x. x ∈ s ⟹ (f n has_field_derivative f' n x) (at x)"
assumes "uniformly_convergent_on s (λn x. ∑i<n. f' i x)"
assumes "x0 ∈ s" "summable (λn. f n x0)" and x: "x ∈ s"
shows "summable (λn. f n x)" and "(λx. ∑n. f n x) field_differentiable (at x)"
proof -
from assms(4) obtain g' where A: "uniform_limit s (λn x. ∑i<n. f' i x) g' sequentially"
unfolding uniformly_convergent_on_def by blast
from x and ‹open s› have s: "at x within s = at x" by (rule at_within_open)
have "∃g. ∀x∈s. (λn. f n x) sums g x ∧ (g has_field_derivative g' x) (at x within s)"
by (intro has_field_derivative_series[of s f f' g' x0] assms A has_field_derivative_at_within)
then obtain g where g: "⋀x. x ∈ s ⟹ (λn. f n x) sums g x"
"⋀x. x ∈ s ⟹ (g has_field_derivative g' x) (at x within s)" by blast
from g[OF x] show "summable (λn. f n x)" by (auto simp: summable_def)
from g(2)[OF x] have g': "(g has_derivative op * (g' x)) (at x)"
by (simp add: has_field_derivative_def s)
have "((λx. ∑n. f n x) has_derivative op * (g' x)) (at x)"
by (rule has_derivative_transform_within_open[OF g' ‹open s› x])
(insert g, auto simp: sums_iff)
thus "(λx. ∑n. f n x) field_differentiable (at x)" unfolding differentiable_def
by (auto simp: summable_def field_differentiable_def has_field_derivative_def)
qed
lemma field_differentiable_series':
fixes f :: "nat ⇒ complex ⇒ complex"
assumes "convex s" "open s"
assumes "⋀n x. x ∈ s ⟹ (f n has_field_derivative f' n x) (at x)"
assumes "uniformly_convergent_on s (λn x. ∑i<n. f' i x)"
assumes "x0 ∈ s" "summable (λn. f n x0)"
shows "(λx. ∑n. f n x) field_differentiable (at x0)"
using field_differentiable_series[OF assms, of x0] ‹x0 ∈ s› by blast+
subsection‹Bound theorem›
lemma field_differentiable_bound:
fixes s :: "complex set"
assumes cvs: "convex s"
and df: "⋀z. z ∈ s ⟹ (f has_field_derivative f' z) (at z within s)"
and dn: "⋀z. z ∈ s ⟹ norm (f' z) ≤ B"
and "x ∈ s" "y ∈ s"
shows "norm(f x - f y) ≤ B * norm(x - y)"
apply (rule differentiable_bound [OF cvs])
apply (rule ballI, erule df [unfolded has_field_derivative_def])
apply (rule ballI, rule onorm_le, simp add: norm_mult mult_right_mono dn)
apply fact
apply fact
done
subsection‹Inverse function theorem for complex derivatives›
lemma has_complex_derivative_inverse_basic:
fixes f :: "complex ⇒ complex"
shows "DERIV f (g y) :> f' ⟹
f' ≠ 0 ⟹
continuous (at y) g ⟹
open t ⟹
y ∈ t ⟹
(⋀z. z ∈ t ⟹ f (g z) = z)
⟹ DERIV g y :> inverse (f')"
unfolding has_field_derivative_def
apply (rule has_derivative_inverse_basic)
apply (auto simp: bounded_linear_mult_right)
done
lemma has_complex_derivative_inverse_strong:
fixes f :: "complex ⇒ complex"
shows "DERIV f x :> f' ⟹
f' ≠ 0 ⟹
open s ⟹
x ∈ s ⟹
continuous_on s f ⟹
(⋀z. z ∈ s ⟹ g (f z) = z)
⟹ DERIV g (f x) :> inverse (f')"
unfolding has_field_derivative_def
apply (rule has_derivative_inverse_strong [of s x f g ])
using assms
by auto
lemma has_complex_derivative_inverse_strong_x:
fixes f :: "complex ⇒ complex"
shows "DERIV f (g y) :> f' ⟹
f' ≠ 0 ⟹
open s ⟹
continuous_on s f ⟹
g y ∈ s ⟹ f(g y) = y ⟹
(⋀z. z ∈ s ⟹ g (f z) = z)
⟹ DERIV g y :> inverse (f')"
unfolding has_field_derivative_def
apply (rule has_derivative_inverse_strong_x [of s g y f])
using assms
by auto
subsection ‹Taylor on Complex Numbers›
lemma setsum_Suc_reindex:
fixes f :: "nat ⇒ 'a::ab_group_add"
shows "setsum f {0..n} = f 0 - f (Suc n) + setsum (λi. f (Suc i)) {0..n}"
by (induct n) auto
lemma complex_taylor:
assumes s: "convex s"
and f: "⋀i x. x ∈ s ⟹ i ≤ n ⟹ (f i has_field_derivative f (Suc i) x) (at x within s)"
and B: "⋀x. x ∈ s ⟹ cmod (f (Suc n) x) ≤ B"
and w: "w ∈ s"
and z: "z ∈ s"
shows "cmod(f 0 z - (∑i≤n. f i w * (z-w) ^ i / (fact i)))
≤ B * cmod(z - w)^(Suc n) / fact n"
proof -
have wzs: "closed_segment w z ⊆ s" using assms
by (metis convex_contains_segment)
{ fix u
assume "u ∈ closed_segment w z"
then have "u ∈ s"
by (metis wzs subsetD)
have "(∑i≤n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) +
f (Suc i) u * (z-u)^i / (fact i)) =
f (Suc n) u * (z-u) ^ n / (fact n)"
proof (induction n)
case 0 show ?case by simp
next
case (Suc n)
have "(∑i≤Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / (fact i) +
f (Suc i) u * (z-u) ^ i / (fact i)) =
f (Suc n) u * (z-u) ^ n / (fact n) +
f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / (fact (Suc n)) -
f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / (fact (Suc n))"
using Suc by simp
also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n))"
proof -
have "(fact(Suc n)) *
(f(Suc n) u *(z-u) ^ n / (fact n) +
f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / (fact(Suc n)) -
f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / (fact(Suc n))) =
((fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / (fact n) +
((fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / (fact(Suc n))) -
((fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / (fact(Suc n))"
by (simp add: algebra_simps del: fact.simps)
also have "... = ((fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / (fact n) +
(f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -
(f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"
by (simp del: fact.simps)
also have "... = (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) +
(f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) -
(f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))"
by (simp only: fact.simps of_nat_mult ac_simps) simp
also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)"
by (simp add: algebra_simps)
finally show ?thesis
by (simp add: mult_left_cancel [where c = "(fact (Suc n))", THEN iffD1] del: fact.simps)
qed
finally show ?case .
qed
then have "((λv. (∑i≤n. f i v * (z - v)^i / (fact i)))
has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n))
(at u within s)"
apply (intro derivative_eq_intros)
apply (blast intro: assms ‹u ∈ s›)
apply (rule refl)+
apply (auto simp: field_simps)
done
} note sum_deriv = this
{ fix u
assume u: "u ∈ closed_segment w z"
then have us: "u ∈ s"
by (metis wzs subsetD)
have "cmod (f (Suc n) u) * cmod (z - u) ^ n ≤ cmod (f (Suc n) u) * cmod (u - z) ^ n"
by (metis norm_minus_commute order_refl)
also have "... ≤ cmod (f (Suc n) u) * cmod (z - w) ^ n"
by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u])
also have "... ≤ B * cmod (z - w) ^ n"
by (metis norm_ge_zero zero_le_power mult_right_mono B [OF us])
finally have "cmod (f (Suc n) u) * cmod (z - u) ^ n ≤ B * cmod (z - w) ^ n" .
} note cmod_bound = this
have "(∑i≤n. f i z * (z - z) ^ i / (fact i)) = (∑i≤n. (f i z / (fact i)) * 0 ^ i)"
by simp
also have "… = f 0 z / (fact 0)"
by (subst setsum_zero_power) simp
finally have "cmod (f 0 z - (∑i≤n. f i w * (z - w) ^ i / (fact i)))
≤ cmod ((∑i≤n. f i w * (z - w) ^ i / (fact i)) -
(∑i≤n. f i z * (z - z) ^ i / (fact i)))"
by (simp add: norm_minus_commute)
also have "... ≤ B * cmod (z - w) ^ n / (fact n) * cmod (w - z)"
apply (rule field_differentiable_bound
[where f' = "λw. f (Suc n) w * (z - w)^n / (fact n)"
and s = "closed_segment w z", OF convex_closed_segment])
apply (auto simp: ends_in_segment DERIV_subset [OF sum_deriv wzs]
norm_divide norm_mult norm_power divide_le_cancel cmod_bound)
done
also have "... ≤ B * cmod (z - w) ^ Suc n / (fact n)"
by (simp add: algebra_simps norm_minus_commute)
finally show ?thesis .
qed
text‹Something more like the traditional MVT for real components›
lemma complex_mvt_line:
assumes "⋀u. u ∈ closed_segment w z ⟹ (f has_field_derivative f'(u)) (at u)"
shows "∃u. u ∈ closed_segment w z ∧ Re(f z) - Re(f w) = Re(f'(u) * (z - w))"
proof -
have twz: "⋀t. (1 - t) *⇩R w + t *⇩R z = w + t *⇩R (z - w)"
by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib)
note assms[unfolded has_field_derivative_def, derivative_intros]
show ?thesis
apply (cut_tac mvt_simple
[of 0 1 "Re o f o (λt. (1 - t) *⇩R w + t *⇩R z)"
"λu. Re o (λh. f'((1 - u) *⇩R w + u *⇩R z) * h) o (λt. t *⇩R (z - w))"])
apply auto
apply (rule_tac x="(1 - x) *⇩R w + x *⇩R z" in exI)
apply (auto simp: closed_segment_def twz) []
apply (intro derivative_eq_intros has_derivative_at_within, simp_all)
apply (simp add: fun_eq_iff real_vector.scale_right_diff_distrib)
apply (force simp: twz closed_segment_def)
done
qed
lemma complex_taylor_mvt:
assumes "⋀i x. ⟦x ∈ closed_segment w z; i ≤ n⟧ ⟹ ((f i) has_field_derivative f (Suc i) x) (at x)"
shows "∃u. u ∈ closed_segment w z ∧
Re (f 0 z) =
Re ((∑i = 0..n. f i w * (z - w) ^ i / (fact i)) +
(f (Suc n) u * (z-u)^n / (fact n)) * (z - w))"
proof -
{ fix u
assume u: "u ∈ closed_segment w z"
have "(∑i = 0..n.
(f (Suc i) u * (z-u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) /
(fact i)) =
f (Suc 0) u -
(f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) /
(fact (Suc n)) +
(∑i = 0..n.
(f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) /
(fact (Suc i)))"
by (subst setsum_Suc_reindex) simp
also have "... = f (Suc 0) u -
(f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) /
(fact (Suc n)) +
(∑i = 0..n.
f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i)) -
f (Suc i) u * (z-u) ^ i / (fact i))"
by (simp only: diff_divide_distrib fact_cancel ac_simps)
also have "... = f (Suc 0) u -
(f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) /
(fact (Suc n)) +
f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n)) - f (Suc 0) u"
by (subst setsum_Suc_diff) auto
also have "... = f (Suc n) u * (z-u) ^ n / (fact n)"
by (simp only: algebra_simps diff_divide_distrib fact_cancel)
finally have "(∑i = 0..n. (f (Suc i) u * (z - u) ^ i
- of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) =
f (Suc n) u * (z - u) ^ n / (fact n)" .
then have "((λu. ∑i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative
f (Suc n) u * (z - u) ^ n / (fact n)) (at u)"
apply (intro derivative_eq_intros)+
apply (force intro: u assms)
apply (rule refl)+
apply (auto simp: ac_simps)
done
}
then show ?thesis
apply (cut_tac complex_mvt_line [of w z "λu. ∑i = 0..n. f i u * (z-u) ^ i / (fact i)"
"λu. (f (Suc n) u * (z-u)^n / (fact n))"])
apply (auto simp add: intro: open_closed_segment)
done
qed
subsection ‹Polynomal function extremal theorem, from HOL Light›
lemma polyfun_extremal_lemma:
fixes c :: "nat ⇒ 'a::real_normed_div_algebra"
assumes "0 < e"
shows "∃M. ∀z. M ≤ norm(z) ⟶ norm (∑i≤n. c(i) * z^i) ≤ e * norm(z) ^ (Suc n)"
proof (induct n)
case 0 with assms
show ?case
apply (rule_tac x="norm (c 0) / e" in exI)
apply (auto simp: field_simps)
done
next
case (Suc n)
obtain M where M: "⋀z. M ≤ norm z ⟹ norm (∑i≤n. c i * z^i) ≤ e * norm z ^ Suc n"
using Suc assms by blast
show ?case
proof (rule exI [where x= "max M (1 + norm(c(Suc n)) / e)"], clarsimp simp del: power_Suc)
fix z::'a
assume z1: "M ≤ norm z" and "1 + norm (c (Suc n)) / e ≤ norm z"
then have z2: "e + norm (c (Suc n)) ≤ e * norm z"
using assms by (simp add: field_simps)
have "norm (∑i≤n. c i * z^i) ≤ e * norm z ^ Suc n"
using M [OF z1] by simp
then have "norm (∑i≤n. c i * z^i) + norm (c (Suc n) * z ^ Suc n) ≤ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)"
by simp
then have "norm ((∑i≤n. c i * z^i) + c (Suc n) * z ^ Suc n) ≤ e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)"
by (blast intro: norm_triangle_le elim: )
also have "... ≤ (e + norm (c (Suc n))) * norm z ^ Suc n"
by (simp add: norm_power norm_mult algebra_simps)
also have "... ≤ (e * norm z) * norm z ^ Suc n"
by (metis z2 mult.commute mult_left_mono norm_ge_zero norm_power)
finally show "norm ((∑i≤n. c i * z^i) + c (Suc n) * z ^ Suc n) ≤ e * norm z ^ Suc (Suc n)"
by simp
qed
qed
lemma polyfun_extremal:
fixes c :: "nat ⇒ 'a::real_normed_div_algebra"
assumes k: "c k ≠ 0" "1≤k" and kn: "k≤n"
shows "eventually (λz. norm (∑i≤n. c(i) * z^i) ≥ B) at_infinity"
using kn
proof (induction n)
case 0
then show ?case
using k by simp
next
case (Suc m)
let ?even = ?case
show ?even
proof (cases "c (Suc m) = 0")
case True
then show ?even using Suc k
by auto (metis antisym_conv less_eq_Suc_le not_le)
next
case False
then obtain M where M:
"⋀z. M ≤ norm z ⟹ norm (∑i≤m. c i * z^i) ≤ norm (c (Suc m)) / 2 * norm z ^ Suc m"
using polyfun_extremal_lemma [of "norm(c (Suc m)) / 2" c m] Suc
by auto
have "∃b. ∀z. b ≤ norm z ⟶ B ≤ norm (∑i≤Suc m. c i * z^i)"
proof (rule exI [where x="max M (max 1 (¦B¦ / (norm(c (Suc m)) / 2)))"], clarsimp simp del: power_Suc)
fix z::'a
assume z1: "M ≤ norm z" "1 ≤ norm z"
and "¦B¦ * 2 / norm (c (Suc m)) ≤ norm z"
then have z2: "¦B¦ ≤ norm (c (Suc m)) * norm z / 2"
using False by (simp add: field_simps)
have nz: "norm z ≤ norm z ^ Suc m"
by (metis ‹1 ≤ norm z› One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc)
have *: "⋀y x. norm (c (Suc m)) * norm z / 2 ≤ norm y - norm x ⟹ B ≤ norm (x + y)"
by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2)
have "norm z * norm (c (Suc m)) + 2 * norm (∑i≤m. c i * z^i)
≤ norm (c (Suc m)) * norm z + norm (c (Suc m)) * norm z ^ Suc m"
using M [of z] Suc z1 by auto
also have "... ≤ 2 * (norm (c (Suc m)) * norm z ^ Suc m)"
using nz by (simp add: mult_mono del: power_Suc)
finally show "B ≤ norm ((∑i≤m. c i * z^i) + c (Suc m) * z ^ Suc m)"
using Suc.IH
apply (auto simp: eventually_at_infinity)
apply (rule *)
apply (simp add: field_simps norm_mult norm_power)
done
qed
then show ?even
by (simp add: eventually_at_infinity)
qed
qed
end