# HG changeset patch # User paulson # Date 1526939536 -3600 # Node ID 009f783d1bac3a2ae736f9cff6f368c12f8ef239 # Parent ddf1ead7b182b3ba88d5c25dc4af71baa277c8bb small clean-up of Complex_Analysis_Basics diff -r ddf1ead7b182 -r 009f783d1bac src/HOL/Analysis/Cauchy_Integral_Theorem.thy --- a/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon May 21 18:36:30 2018 +0200 +++ b/src/HOL/Analysis/Cauchy_Integral_Theorem.thy Mon May 21 22:52:16 2018 +0100 @@ -5936,9 +5936,9 @@ also have "... = u^n * deriv (\z. (deriv ^^ n) f (u * z)) z" apply (rule deriv_cmult) apply (rule analytic_on_imp_differentiable_at [OF _ Suc.prems]) - apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and t=t, unfolded o_def]) - apply (simp add: analytic_on_linear) - apply (simp add: analytic_on_open f holomorphic_higher_deriv t) + apply (rule analytic_on_compose_gen [where g="(deriv ^^ n) f" and T=t, unfolded o_def]) + apply (simp add: analytic_on_linear) + apply (simp add: analytic_on_open f holomorphic_higher_deriv t) apply (blast intro: fg) done also have "... = u * u ^ n * deriv ((deriv ^^ n) f) (u * z)" diff -r ddf1ead7b182 -r 009f783d1bac src/HOL/Analysis/Complex_Analysis_Basics.thy --- a/src/HOL/Analysis/Complex_Analysis_Basics.thy Mon May 21 18:36:30 2018 +0200 +++ b/src/HOL/Analysis/Complex_Analysis_Basics.thy Mon May 21 22:52:16 2018 +0100 @@ -32,7 +32,7 @@ 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) + using of_nat_neq_0 by force lemma bilinear_times: fixes c::"'a::real_algebra" shows "bilinear (\x y::'a. x*y)" @@ -41,34 +41,6 @@ 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) = ( * ) 0" by auto @@ -116,48 +88,48 @@ 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" + 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) lemmas DERIV_zero_constant = has_field_derivative_zero_constant lemma DERIV_zero_unique: - assumes "convex s" - and d0: "\x. x\s \ (f has_field_derivative 0) (at x within s)" - and "a \ s" - and "x \ s" + 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: - assumes "connected s" - and "open s" - and d0: "\x. x\s \ DERIV f x :> 0" - and "a \ s" - and "x \ s" + 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)" + 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" + 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) @@ -270,8 +242,6 @@ subsection\Holomorphic functions\ -subsection\Holomorphic functions\ - 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)" @@ -455,20 +425,29 @@ unfolding DERIV_deriv_iff_field_differentiable[symmetric] by (auto intro!: DERIV_imp_deriv derivative_eq_intros) -lemma deriv_cmult [simp]: +lemma deriv_cmult: "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) + by simp -lemma deriv_cmult_right [simp]: +lemma deriv_cmult_right: "f field_differentiable at z \ deriv (\w. f w * c) z = deriv f z * c" + by simp + +lemma deriv_inverse [simp]: + "\f field_differentiable at z; f z \ 0\ + \ deriv (\w. inverse (f w)) z = - deriv f z / f z ^ 2" unfolding DERIV_deriv_iff_field_differentiable[symmetric] - by (auto intro!: DERIV_imp_deriv derivative_eq_intros) + by (safe intro!: DERIV_imp_deriv derivative_eq_intros) (auto simp: divide_simps power2_eq_square) -lemma deriv_cdivide_right [simp]: +lemma deriv_divide [simp]: + "\f field_differentiable at z; g field_differentiable at z; g z \ 0\ + \ deriv (\w. f w / g w) z = (deriv f z * g z - f z * deriv g z) / g z ^ 2" + by (simp add: field_class.field_divide_inverse field_differentiable_inverse) + (simp add: divide_simps power2_eq_square) + +lemma deriv_cdivide_right: "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) + by (simp add: field_class.field_divide_inverse) 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\ @@ -480,10 +459,9 @@ 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 + unfolding DERIV_deriv_iff_field_differentiable [symmetric] + by (metis (full_types) DERIV_chain2 DERIV_cmult_Id mult.commute) + lemma nonzero_deriv_nonconstant: assumes df: "DERIV f \ :> df" and S: "open S" "\ \ S" and "df \ 0" @@ -494,10 +472,8 @@ 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 + by (rule nonzero_deriv_nonconstant [of f "deriv f \" \ S]) + (use assms in \auto simp: holomorphic_derivI\) subsection\Caratheodory characterization\ @@ -516,53 +492,52 @@ 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)" + where "f analytic_on S \ \x \ S. \e. 0 < e \ f holomorphic_on (ball x e)" named_theorems analytic_intros "introduction rules for proving analyticity" -lemma analytic_imp_holomorphic: "f analytic_on s \ f holomorphic_on s" +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" +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)" + "f analytic_on S \ x \ S \ f field_differentiable (at x)" apply (auto simp: analytic_on_def holomorphic_on_def) by (metis open_ball centre_in_ball field_differentiable_within_open) -lemma analytic_on_subset: "f analytic_on s \ t \ s \ f analytic_on t" +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" +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)" +lemma analytic_on_Union: "f analytic_on (\\) \ (\T \ \. 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))" +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)" + "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)" + 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" + 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 (rule exI [where x="\{U. open U \ f analytic_on U}"], auto) apply (metis 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" + 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" @@ -570,26 +545,26 @@ finally show ?thesis . qed -lemma analytic_on_linear [analytic_intros,simp]: "(( * ) c) analytic_on s" +lemma analytic_on_linear [analytic_intros,simp]: "(( * ) c) analytic_on S" by (auto simp add: analytic_on_holomorphic) -lemma analytic_on_const [analytic_intros,simp]: "(\z. c) analytic_on s" +lemma analytic_on_const [analytic_intros,simp]: "(\z. c) analytic_on S" by (metis analytic_on_def holomorphic_on_const zero_less_one) -lemma analytic_on_ident [analytic_intros,simp]: "(\x. x) analytic_on s" +lemma analytic_on_ident [analytic_intros,simp]: "(\x. x) analytic_on S" by (simp add: analytic_on_def gt_ex) -lemma analytic_on_id [analytic_intros]: "id analytic_on s" +lemma analytic_on_id [analytic_intros]: "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" + 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" + 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 @@ -607,22 +582,22 @@ 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" + "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 [analytic_intros]: - "f analytic_on s \ (\z. -(f z)) analytic_on s" + "f analytic_on S \ (\z. -(f z)) analytic_on S" by (metis analytic_on_holomorphic holomorphic_on_minus) lemma analytic_on_add [analytic_intros]: - assumes f: "f analytic_on s" - and g: "g analytic_on s" - shows "(\z. f z + g z) analytic_on s" + 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" + 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 @@ -636,13 +611,13 @@ qed lemma analytic_on_diff [analytic_intros]: - assumes f: "f analytic_on s" - and g: "g analytic_on s" - shows "(\z. f z - g z) analytic_on s" + 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" + 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 @@ -656,13 +631,13 @@ qed lemma analytic_on_mult [analytic_intros]: - assumes f: "f analytic_on s" - and g: "g analytic_on s" - shows "(\z. f z * g z) analytic_on s" + 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" + 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 @@ -676,13 +651,13 @@ qed lemma analytic_on_inverse [analytic_intros]: - assumes f: "f analytic_on s" - and nz: "(\z. z \ s \ f z \ 0)" - shows "(\z. inverse (f z)) analytic_on s" + 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" + 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" @@ -698,19 +673,19 @@ qed lemma analytic_on_divide [analytic_intros]: - 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" + 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 [analytic_intros]: - "f analytic_on s \ (\z. (f z) ^ n) analytic_on s" + "f analytic_on S \ (\z. (f z) ^ n) analytic_on S" by (induct n) (auto simp: analytic_on_mult) lemma analytic_on_sum [analytic_intros]: - "(\i. i \ I \ (f i) analytic_on s) \ (\x. sum (\i. f i x) I) analytic_on s" + "(\i. i \ I \ (f i) analytic_on S) \ (\x. sum (\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: @@ -727,10 +702,10 @@ 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 + proof (rule complex_derivative_transform_within_open [where s=S]) + show "g \ f holomorphic_on S" + by (rule assms holomorphic_on_compose_gen holomorphic_intros)+ + qed (use assms in auto) also have "... = 1" by simp finally show ?thesis . @@ -811,23 +786,23 @@ (* TODO: Could probably be simplified using Uniform_Limit *) 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)" + 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" + 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" + 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" + 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" + 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" @@ -841,30 +816,30 @@ proof (rule has_derivative_sequence [OF cvs _ _ x]) show "(\n. f n x) \ l" by (rule tf) - next show "\e. e > 0 \ \\<^sub>F n in sequentially. \x\s. \h. cmod (f' n x * h - g' x * h) \ e * cmod h" + next show "\e. e > 0 \ \\<^sub>F n in sequentially. \x\S. \h. cmod (f' n x * h - g' x * h) \ e * cmod h" unfolding eventually_sequentially by (blast intro: **) qed (metis has_field_derivative_def df) 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 + 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 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))" + 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)" + 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 + then obtain N where N: "\n x. n \ N \ x \ S \ cmod ((\i e" by (metis conv) - have "\N. \n\N. \x\s. \h. cmod ((\i e * cmod h" + have "\N. \n\N. \x\S. \h. cmod ((\i e * cmod h" proof (rule exI [of _ N], clarify) fix n y h - assume "N \ n" "y \ s" + assume "N \ n" "y \ S" then have "cmod ((\i e" by (metis N) then have "cmod h * cmod ((\i cmod h * e" @@ -877,12 +852,12 @@ 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)" + 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. e>0 \ \\<^sub>F n in sequentially. \x\s. \h. cmod ((\i e * cmod h" + next show "\e. e>0 \ \\<^sub>F n in sequentially. \x\S. \h. cmod ((\i e * cmod h" unfolding eventually_sequentially by (blast intro: **) qed qed @@ -890,23 +865,23 @@ lemma field_differentiable_series: fixes f :: "nat \ 'a::{real_normed_field,banach} \ 'a" - 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 s" "summable (\n. f n x0)" and x: "x \ s" + 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 S" "summable (\n. f n x0)" and x: "x \ S" shows "(\x. \n. f n x) field_differentiable (at x)" proof - - from assms(4) obtain g' where A: "uniform_limit s (\n x. \in x. \iopen 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 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(2)[OF x] have g': "(g has_derivative ( * ) (g' x)) (at x)" - by (simp add: has_field_derivative_def s) + by (simp add: has_field_derivative_def S) have "((\x. \n. f n x) has_derivative ( * ) (g' x)) (at x)" - by (rule has_derivative_transform_within_open[OF g' \open s\ 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) @@ -915,11 +890,11 @@ subsection\Bound theorem\ lemma field_differentiable_bound: - fixes s :: "'a::real_normed_field 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" + fixes S :: "'a::real_normed_field 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 (erule df [unfolded has_field_derivative_def]) @@ -941,35 +916,31 @@ apply (auto simp: bounded_linear_mult_right) done -lemmas has_complex_derivative_inverse_basic = has_field_derivative_inverse_basic - lemma has_field_derivative_inverse_strong: fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" shows "DERIV f x :> f' \ f' \ 0 \ - open s \ - x \ s \ - continuous_on s f \ - (\z. z \ s \ g (f z) = z) + 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 ]) + apply (rule has_derivative_inverse_strong [of S x f g ]) by auto -lemmas has_complex_derivative_inverse_strong = has_field_derivative_inverse_strong lemma has_field_derivative_inverse_strong_x: fixes f :: "'a::{euclidean_space,real_normed_field} \ 'a" 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) + 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]) + apply (rule has_derivative_inverse_strong_x [of S g y f]) by auto -lemmas has_complex_derivative_inverse_strong_x = has_field_derivative_inverse_strong_x subsection \Taylor on Complex Numbers\ @@ -979,19 +950,19 @@ by (induct n) auto lemma field_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 \ norm (f (Suc n) x) \ B" - and w: "w \ s" - and z: "z \ s" + 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 \ norm (f (Suc n) x) \ B" + and w: "w \ S" + and z: "z \ S" shows "norm(f 0 z - (\i\n. f i w * (z-w) ^ i / (fact i))) \ B * norm(z - w)^(Suc n) / fact n" proof - - have wzs: "closed_segment w z \ s" using assms + 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" + 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)) = @@ -1033,16 +1004,16 @@ 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)" + (at u within S)" apply (intro derivative_eq_intros) - apply (blast intro: assms \u \ s\) + 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" + then have us: "u \ S" by (metis wzs subsetD) have "norm (f (Suc n) u) * norm (z - u) ^ n \ norm (f (Suc n) u) * norm (u - z) ^ n" by (metis norm_minus_commute order_refl) @@ -1063,7 +1034,7 @@ also have "... \ B * norm (z - w) ^ n / (fact n) * norm (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]) + 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 @@ -1073,11 +1044,11 @@ qed 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" + 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" using assms by (rule field_taylor) diff -r ddf1ead7b182 -r 009f783d1bac src/HOL/Analysis/Complex_Transcendental.thy --- a/src/HOL/Analysis/Complex_Transcendental.thy Mon May 21 18:36:30 2018 +0200 +++ b/src/HOL/Analysis/Complex_Transcendental.thy Mon May 21 22:52:16 2018 +0100 @@ -1234,10 +1234,10 @@ have "(exp has_field_derivative z) (at (Ln z))" by (metis znz DERIV_exp exp_Ln) then show "(Ln has_field_derivative inverse(z)) (at z)" - apply (rule has_complex_derivative_inverse_strong_x - [where s = "{w. -pi < Im(w) \ Im(w) < pi}"]) + apply (rule has_field_derivative_inverse_strong_x + [where S = "{w. -pi < Im(w) \ Im(w) < pi}"]) using znz * - apply (auto simp: Transcendental.continuous_on_exp [OF continuous_on_id] open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt mpi_less_Im_Ln) + apply (auto simp: continuous_on_exp [OF continuous_on_id] open_Collect_conj open_halfspace_Im_gt open_halfspace_Im_lt mpi_less_Im_Ln) done qed @@ -3054,7 +3054,7 @@ then have "cos (Arcsin z) \ 0" by (metis diff_0_right power_zero_numeral sin_squared_eq) then show ?thesis - apply (rule has_complex_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) + apply (rule has_field_derivative_inverse_basic [OF DERIV_sin _ _ open_ball [of z 1]]) apply (auto intro: isCont_Arcsin assms) done qed @@ -3219,7 +3219,7 @@ then have "- sin (Arccos z) \ 0" by (metis cos_squared_eq diff_0_right mult_zero_left neg_0_equal_iff_equal power2_eq_square) then have "(Arccos has_field_derivative inverse(- sin(Arccos z))) (at z)" - apply (rule has_complex_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]]) + apply (rule has_field_derivative_inverse_basic [OF DERIV_cos _ _ open_ball [of z 1]]) apply (auto intro: isCont_Arccos assms) done then show ?thesis diff -r ddf1ead7b182 -r 009f783d1bac src/HOL/Analysis/Conformal_Mappings.thy --- a/src/HOL/Analysis/Conformal_Mappings.thy Mon May 21 18:36:30 2018 +0200 +++ b/src/HOL/Analysis/Conformal_Mappings.thy Mon May 21 22:52:16 2018 +0100 @@ -1290,7 +1290,7 @@ have 2: "deriv f z \ 0" using \z \ S\ \open S\ holf holomorphic_injective_imp_regular injf by blast show ?thesis - apply (rule has_complex_derivative_inverse_strong [OF 1 2 \open S\ \z \ S\]) + apply (rule has_field_derivative_inverse_strong [OF 1 2 \open S\ \z \ S\]) apply (simp add: holf holomorphic_on_imp_continuous_on) by (simp add: injf the_inv_into_f_f) qed diff -r ddf1ead7b182 -r 009f783d1bac src/HOL/Analysis/Great_Picard.thy --- a/src/HOL/Analysis/Great_Picard.thy Mon May 21 18:36:30 2018 +0200 +++ b/src/HOL/Analysis/Great_Picard.thy Mon May 21 22:52:16 2018 +0100 @@ -1136,7 +1136,7 @@ apply (metis Suc_pred mult.commute power_Suc) done then show ?thesis - apply (rule DERIV_imp_deriv [OF DERIV_transform_within_open [where s = "ball z0 r"]]) + apply (rule DERIV_imp_deriv [OF DERIV_transform_within_open [where S = "ball z0 r"]]) using that \m > 0\ \0 < r\ apply (simp_all add: hnz geq) done