| author | paulson <lp15@cam.ac.uk> | 
| Tue, 08 Aug 2017 23:54:49 +0200 | |
| changeset 66384 | cc66710c9d48 | 
| parent 66252 | b73f94b366b7 | 
| child 66453 | cc19f7ca2ed6 | 
| 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 | ||
| 60420 | 5 | section \<open>Complex Analysis Basics\<close> | 
| 56215 | 6 | |
| 7 | theory Complex_Analysis_Basics | |
| 63941 
f353674c2528
move absolutely_integrable_on to Equivalence_Lebesgue_Henstock_Integration, now based on the Lebesgue integral
 hoelzl parents: 
63918diff
changeset | 8 | imports Equivalence_Lebesgue_Henstock_Integration "~~/src/HOL/Library/Nonpos_Ints" | 
| 56215 | 9 | begin | 
| 10 | ||
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 11 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 12 | subsection\<open>General lemmas\<close> | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 13 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 14 | lemma nonneg_Reals_cmod_eq_Re: "z \<in> \<real>\<^sub>\<ge>\<^sub>0 \<Longrightarrow> norm z = Re z" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 15 | by (simp add: complex_nonneg_Reals_iff cmod_eq_Re) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 16 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 17 | lemma has_derivative_mult_right: | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 18 | fixes c:: "'a :: real_normed_algebra" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 19 | shows "((op * c) has_derivative (op * c)) F" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 20 | by (rule has_derivative_mult_right [OF has_derivative_id]) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 21 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 22 | lemma has_derivative_of_real[derivative_intros, simp]: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 23 | "(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 | 24 | using bounded_linear.has_derivative[OF bounded_linear_of_real] . | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 25 | |
| 66252 | 26 | lemma has_vector_derivative_real_field: | 
| 61806 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61610diff
changeset | 27 | "DERIV f (of_real a) :> f' \<Longrightarrow> ((\<lambda>x. f (of_real x)) has_vector_derivative f') (at a within s)" | 
| 
d2e62ae01cd8
Cauchy's integral formula for circles.  Starting to fix eventually_mono.
 paulson <lp15@cam.ac.uk> parents: 
61610diff
changeset | 28 | using has_derivative_compose[of of_real of_real a _ f "op * f'"] | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 29 | by (simp add: scaleR_conv_of_real ac_simps has_vector_derivative_def has_field_derivative_def) | 
| 66252 | 30 | lemmas has_vector_derivative_real_complex = has_vector_derivative_real_field | 
| 56215 | 31 | |
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 32 | lemma fact_cancel: | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 33 | fixes c :: "'a::real_field" | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 34 | shows "of_nat (Suc n) * c / (fact (Suc n)) = c / (fact n)" | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 35 | by (simp add: of_nat_mult del: of_nat_Suc times_nat.simps) | 
| 56889 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56479diff
changeset | 36 | |
| 56215 | 37 | lemma bilinear_times: | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 38 | 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 | 39 | by (auto simp: bilinear_def distrib_left distrib_right intro!: linearI) | 
| 56215 | 40 | |
| 41 | lemma linear_cnj: "linear cnj" | |
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 42 | using bounded_linear.linear[OF bounded_linear_cnj] . | 
| 56215 | 43 | |
| 44 | lemma tendsto_Re_upper: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 45 | assumes "~ (trivial_limit F)" | 
| 61973 | 46 | "(f \<longlongrightarrow> l) F" | 
| 56215 | 47 | "eventually (\<lambda>x. Re(f x) \<le> b) F" | 
| 48 | shows "Re(l) \<le> b" | |
| 49 | by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Re) | |
| 50 | ||
| 51 | lemma tendsto_Re_lower: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 52 | assumes "~ (trivial_limit F)" | 
| 61973 | 53 | "(f \<longlongrightarrow> l) F" | 
| 56215 | 54 | "eventually (\<lambda>x. b \<le> Re(f x)) F" | 
| 55 | shows "b \<le> Re(l)" | |
| 56 | by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Re) | |
| 57 | ||
| 58 | lemma tendsto_Im_upper: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 59 | assumes "~ (trivial_limit F)" | 
| 61973 | 60 | "(f \<longlongrightarrow> l) F" | 
| 56215 | 61 | "eventually (\<lambda>x. Im(f x) \<le> b) F" | 
| 62 | shows "Im(l) \<le> b" | |
| 63 | by (metis assms tendsto_le [OF _ tendsto_const] tendsto_Im) | |
| 64 | ||
| 65 | lemma tendsto_Im_lower: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 66 | assumes "~ (trivial_limit F)" | 
| 61973 | 67 | "(f \<longlongrightarrow> l) F" | 
| 56215 | 68 | "eventually (\<lambda>x. b \<le> Im(f x)) F" | 
| 69 | shows "b \<le> Im(l)" | |
| 70 | by (metis assms tendsto_le [OF _ _ tendsto_const] tendsto_Im) | |
| 71 | ||
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 72 | lemma lambda_zero: "(\<lambda>h::'a::mult_zero. 0) = op * 0" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 73 | by auto | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 74 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 75 | lemma lambda_one: "(\<lambda>x::'a::monoid_mult. x) = op * 1" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 76 | by auto | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 77 | |
| 56215 | 78 | lemma continuous_mult_left: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 79 | fixes c::"'a::real_normed_algebra" | 
| 56215 | 80 | shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. c * f x)" | 
| 81 | by (rule continuous_mult [OF continuous_const]) | |
| 82 | ||
| 83 | lemma continuous_mult_right: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 84 | fixes c::"'a::real_normed_algebra" | 
| 56215 | 85 | shows "continuous F f \<Longrightarrow> continuous F (\<lambda>x. f x * c)" | 
| 86 | by (rule continuous_mult [OF _ continuous_const]) | |
| 87 | ||
| 88 | lemma continuous_on_mult_left: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 89 | fixes c::"'a::real_normed_algebra" | 
| 56215 | 90 | shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. c * f x)" | 
| 91 | by (rule continuous_on_mult [OF continuous_on_const]) | |
| 92 | ||
| 93 | lemma continuous_on_mult_right: | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 94 | fixes c::"'a::real_normed_algebra" | 
| 56215 | 95 | shows "continuous_on s f \<Longrightarrow> continuous_on s (\<lambda>x. f x * c)" | 
| 96 | by (rule continuous_on_mult [OF _ continuous_on_const]) | |
| 97 | ||
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56370diff
changeset | 98 | lemma uniformly_continuous_on_cmul_right [continuous_intros]: | 
| 56215 | 99 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" | 
| 56332 | 100 | shows "uniformly_continuous_on s f \<Longrightarrow> uniformly_continuous_on s (\<lambda>x. f x * c)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 101 | using bounded_linear.uniformly_continuous_on[OF bounded_linear_mult_left] . | 
| 56215 | 102 | |
| 56371 
fb9ae0727548
extend continuous_intros; remove continuous_on_intros and isCont_intros
 hoelzl parents: 
56370diff
changeset | 103 | lemma uniformly_continuous_on_cmul_left[continuous_intros]: | 
| 56215 | 104 | fixes f :: "'a::real_normed_vector \<Rightarrow> 'b::real_normed_algebra" | 
| 105 | assumes "uniformly_continuous_on s f" | |
| 106 | shows "uniformly_continuous_on s (\<lambda>x. c * f x)" | |
| 107 | by (metis assms bounded_linear.uniformly_continuous_on bounded_linear_mult_right) | |
| 108 | ||
| 109 | lemma continuous_within_norm_id [continuous_intros]: "continuous (at x within S) norm" | |
| 110 | by (rule continuous_norm [OF continuous_ident]) | |
| 111 | ||
| 112 | 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 | 113 | by (intro continuous_on_id continuous_on_norm) | 
| 56215 | 114 | |
| 60420 | 115 | subsection\<open>DERIV stuff\<close> | 
| 56215 | 116 | |
| 117 | lemma DERIV_zero_connected_constant: | |
| 118 |   fixes f :: "'a::{real_normed_field,euclidean_space} \<Rightarrow> 'a"
 | |
| 119 | assumes "connected s" | |
| 120 | and "open s" | |
| 121 | and "finite k" | |
| 122 | and "continuous_on s f" | |
| 123 | and "\<forall>x\<in>(s - k). DERIV f x :> 0" | |
| 124 | obtains c where "\<And>x. x \<in> s \<Longrightarrow> f(x) = c" | |
| 125 | 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 | 126 | by (metis DERIV_const has_derivative_const Diff_iff at_within_open frechet_derivative_at has_field_derivative_def) | 
| 56215 | 127 | |
| 66252 | 128 | lemmas DERIV_zero_constant = has_field_derivative_zero_constant | 
| 56215 | 129 | |
| 130 | lemma DERIV_zero_unique: | |
| 131 | assumes "convex s" | |
| 132 | and d0: "\<And>x. x\<in>s \<Longrightarrow> (f has_field_derivative 0) (at x within s)" | |
| 133 | and "a \<in> s" | |
| 134 | and "x \<in> s" | |
| 135 | shows "f x = f a" | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 136 | by (rule has_derivative_zero_unique [OF assms(1) _ assms(4,3)]) | 
| 56332 | 137 | (metis d0 has_field_derivative_imp_has_derivative lambda_zero) | 
| 56215 | 138 | |
| 139 | lemma DERIV_zero_connected_unique: | |
| 140 | assumes "connected s" | |
| 141 | and "open s" | |
| 142 | and d0: "\<And>x. x\<in>s \<Longrightarrow> DERIV f x :> 0" | |
| 143 | and "a \<in> s" | |
| 144 | and "x \<in> s" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 145 | shows "f x = f a" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 146 | 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 | 147 | (metis has_field_derivative_def lambda_zero d0) | 
| 56215 | 148 | |
| 149 | lemma DERIV_transform_within: | |
| 150 | assumes "(f has_field_derivative f') (at a within s)" | |
| 151 | and "0 < d" "a \<in> s" | |
| 152 | and "\<And>x. x\<in>s \<Longrightarrow> dist x a < d \<Longrightarrow> f x = g x" | |
| 153 | shows "(g has_field_derivative f') (at a within s)" | |
| 154 | using assms unfolding has_field_derivative_def | |
| 56332 | 155 | by (blast intro: has_derivative_transform_within) | 
| 56215 | 156 | |
| 157 | lemma DERIV_transform_within_open: | |
| 158 | assumes "DERIV f a :> f'" | |
| 159 | and "open s" "a \<in> s" | |
| 160 | and "\<And>x. x\<in>s \<Longrightarrow> f x = g x" | |
| 161 | shows "DERIV g a :> f'" | |
| 162 | using assms unfolding has_field_derivative_def | |
| 163 | by (metis has_derivative_transform_within_open) | |
| 164 | ||
| 165 | lemma DERIV_transform_at: | |
| 166 | assumes "DERIV f a :> f'" | |
| 167 | and "0 < d" | |
| 168 | and "\<And>x. dist x a < d \<Longrightarrow> f x = g x" | |
| 169 | shows "DERIV g a :> f'" | |
| 170 | by (blast intro: assms DERIV_transform_within) | |
| 171 | ||
| 59615 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
59554diff
changeset | 172 | (*generalising DERIV_isconst_all, which requires type real (using the ordering)*) | 
| 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
59554diff
changeset | 173 | lemma DERIV_zero_UNIV_unique: | 
| 66252 | 174 | "(\<And>x. DERIV f x :> 0) \<Longrightarrow> f x = f a" | 
| 175 | by (metis DERIV_zero_unique UNIV_I convex_UNIV) | |
| 59615 
fdfdf89a83a6
A few new lemmas and a bit of tidying up
 paulson <lp15@cam.ac.uk> parents: 
59554diff
changeset | 176 | |
| 60420 | 177 | subsection \<open>Some limit theorems about real part of real series etc.\<close> | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 178 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 179 | (*MOVE? But not to Finite_Cartesian_Product*) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 180 | lemma sums_vec_nth : | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 181 | assumes "f sums a" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 182 | shows "(\<lambda>x. f x $ i) sums a $ i" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 183 | using assms unfolding sums_def | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 184 | by (auto dest: tendsto_vec_nth [where i=i]) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 185 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 186 | lemma summable_vec_nth : | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 187 | assumes "summable f" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 188 | shows "summable (\<lambda>x. f x $ i)" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 189 | using assms unfolding summable_def | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 190 | by (blast intro: sums_vec_nth) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 191 | |
| 60420 | 192 | subsection \<open>Complex number lemmas\<close> | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 193 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 194 | lemma | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 195 |   shows open_halfspace_Re_lt: "open {z. Re(z) < b}"
 | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 196 |     and open_halfspace_Re_gt: "open {z. Re(z) > b}"
 | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 197 |     and closed_halfspace_Re_ge: "closed {z. Re(z) \<ge> b}"
 | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 198 |     and closed_halfspace_Re_le: "closed {z. Re(z) \<le> b}"
 | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 199 |     and closed_halfspace_Re_eq: "closed {z. Re(z) = b}"
 | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 200 |     and open_halfspace_Im_lt: "open {z. Im(z) < b}"
 | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 201 |     and open_halfspace_Im_gt: "open {z. Im(z) > b}"
 | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 202 |     and closed_halfspace_Im_ge: "closed {z. Im(z) \<ge> b}"
 | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 203 |     and closed_halfspace_Im_le: "closed {z. Im(z) \<le> b}"
 | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 204 |     and closed_halfspace_Im_eq: "closed {z. Im(z) = b}"
 | 
| 63332 | 205 | by (intro open_Collect_less closed_Collect_le closed_Collect_eq continuous_on_Re | 
| 206 | continuous_on_Im continuous_on_id continuous_on_const)+ | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 207 | |
| 61070 | 208 | lemma closed_complex_Reals: "closed (\<real> :: complex set)" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 209 | proof - | 
| 61070 | 210 |   have "(\<real> :: complex set) = {z. Im z = 0}"
 | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 211 | by (auto simp: complex_is_Real_iff) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 212 | then show ?thesis | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 213 | by (metis closed_halfspace_Im_eq) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 214 | qed | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 215 | |
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 216 | lemma closed_Real_halfspace_Re_le: "closed (\<real> \<inter> {w. Re w \<le> x})"
 | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 217 | by (simp add: closed_Int closed_complex_Reals closed_halfspace_Re_le) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 218 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 219 | corollary closed_nonpos_Reals_complex [simp]: "closed (\<real>\<^sub>\<le>\<^sub>0 :: complex set)" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 220 | proof - | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 221 |   have "\<real>\<^sub>\<le>\<^sub>0 = \<real> \<inter> {z. Re(z) \<le> 0}"
 | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 222 | using complex_nonpos_Reals_iff complex_is_Real_iff by auto | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 223 | then show ?thesis | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 224 | by (metis closed_Real_halfspace_Re_le) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 225 | qed | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 226 | |
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 227 | lemma closed_Real_halfspace_Re_ge: "closed (\<real> \<inter> {w. x \<le> Re(w)})"
 | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 228 | using closed_halfspace_Re_ge | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 229 | by (simp add: closed_Int closed_complex_Reals) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 230 | |
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 231 | corollary closed_nonneg_Reals_complex [simp]: "closed (\<real>\<^sub>\<ge>\<^sub>0 :: complex set)" | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 232 | proof - | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 233 |   have "\<real>\<^sub>\<ge>\<^sub>0 = \<real> \<inter> {z. Re(z) \<ge> 0}"
 | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 234 | using complex_nonneg_Reals_iff complex_is_Real_iff by auto | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 235 | then show ?thesis | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 236 | by (metis closed_Real_halfspace_Re_ge) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 237 | qed | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 238 | |
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 239 | lemma closed_real_abs_le: "closed {w \<in> \<real>. \<bar>Re w\<bar> \<le> r}"
 | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 240 | proof - | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 241 |   have "{w \<in> \<real>. \<bar>Re w\<bar> \<le> r} = (\<real> \<inter> {w. Re w \<le> r}) \<inter> (\<real> \<inter> {w. Re w \<ge> -r})"
 | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 242 | by auto | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 243 |   then show "closed {w \<in> \<real>. \<bar>Re w\<bar> \<le> r}"
 | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 244 | by (simp add: closed_Int closed_Real_halfspace_Re_ge closed_Real_halfspace_Re_le) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 245 | qed | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 246 | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 247 | lemma real_lim: | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 248 | fixes l::complex | 
| 61973 | 249 | assumes "(f \<longlongrightarrow> l) F" and "~(trivial_limit F)" and "eventually P F" and "\<And>a. P a \<Longrightarrow> f a \<in> \<real>" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 250 | shows "l \<in> \<real>" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 251 | 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 | 252 | show "eventually (\<lambda>x. f x \<in> \<real>) F" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 253 | using assms(3, 4) by (auto intro: eventually_mono) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 254 | qed | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 255 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 256 | lemma real_lim_sequentially: | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 257 | fixes l::complex | 
| 61973 | 258 | shows "(f \<longlongrightarrow> l) sequentially \<Longrightarrow> (\<exists>N. \<forall>n\<ge>N. f n \<in> \<real>) \<Longrightarrow> l \<in> \<real>" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 259 | by (rule real_lim [where F=sequentially]) (auto simp: eventually_sequentially) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 260 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 261 | lemma real_series: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 262 | fixes l::complex | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 263 | 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 | 264 | unfolding sums_def | 
| 64267 | 265 | by (metis real_lim_sequentially sum_in_Reals) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 266 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 267 | lemma Lim_null_comparison_Re: | 
| 61973 | 268 | assumes "eventually (\<lambda>x. norm(f x) \<le> Re(g x)) F" "(g \<longlongrightarrow> 0) F" shows "(f \<longlongrightarrow> 0) F" | 
| 56889 
48a745e1bde7
avoid the Complex constructor, use the more natural Re/Im view; moved csqrt to Complex.
 hoelzl parents: 
56479diff
changeset | 269 | by (rule Lim_null_comparison[OF assms(1)] tendsto_eq_intros assms(2))+ simp | 
| 56215 | 270 | |
| 60420 | 271 | subsection\<open>Holomorphic functions\<close> | 
| 56215 | 272 | |
| 64394 | 273 | subsection\<open>Holomorphic functions\<close> | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 274 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 275 | definition holomorphic_on :: "[complex \<Rightarrow> complex, complex set] \<Rightarrow> bool" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 276 | (infixl "(holomorphic'_on)" 50) | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 277 | where "f holomorphic_on s \<equiv> \<forall>x\<in>s. f field_differentiable (at x within s)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 278 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 279 | named_theorems holomorphic_intros "structural introduction rules for holomorphic_on" | 
| 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 280 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 281 | lemma holomorphic_onI [intro?]: "(\<And>x. x \<in> s \<Longrightarrow> f field_differentiable (at x within s)) \<Longrightarrow> f holomorphic_on s" | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 282 | by (simp add: holomorphic_on_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 283 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 284 | lemma holomorphic_onD [dest?]: "\<lbrakk>f holomorphic_on s; x \<in> s\<rbrakk> \<Longrightarrow> f field_differentiable (at x within s)" | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 285 | by (simp add: holomorphic_on_def) | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 286 | |
| 64394 | 287 | lemma holomorphic_on_imp_differentiable_on: | 
| 288 | "f holomorphic_on s \<Longrightarrow> f differentiable_on s" | |
| 289 | unfolding holomorphic_on_def differentiable_on_def | |
| 290 | by (simp add: field_differentiable_imp_differentiable) | |
| 291 | ||
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 292 | lemma holomorphic_on_imp_differentiable_at: | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 293 | "\<lbrakk>f holomorphic_on s; open s; x \<in> s\<rbrakk> \<Longrightarrow> f field_differentiable (at x)" | 
| 62131 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 294 | using at_within_open holomorphic_on_def by fastforce | 
| 
1baed43f453e
nonneg_Reals, nonpos_Reals, Cauchy integral formula, etc.
 paulson parents: 
62087diff
changeset | 295 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 296 | lemma holomorphic_on_empty [holomorphic_intros]: "f holomorphic_on {}"
 | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 297 | by (simp add: holomorphic_on_def) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 298 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 299 | lemma holomorphic_on_open: | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 300 | "open s \<Longrightarrow> f holomorphic_on s \<longleftrightarrow> (\<forall>x \<in> s. \<exists>f'. DERIV f x :> f')" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 301 | by (auto simp: holomorphic_on_def field_differentiable_def has_field_derivative_def at_within_open [of _ s]) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 302 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 303 | lemma holomorphic_on_imp_continuous_on: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 304 | "f holomorphic_on s \<Longrightarrow> continuous_on s f" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 305 | by (metis field_differentiable_imp_continuous_at continuous_on_eq_continuous_within holomorphic_on_def) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 306 | |
| 62540 
f2fc5485e3b0
Wenda Li's new material: residue theorem, argument_principle, Rouche_theorem
 paulson <lp15@cam.ac.uk> parents: 
62534diff
changeset | 307 | lemma holomorphic_on_subset [elim]: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 308 | "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 | 309 | unfolding holomorphic_on_def | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 310 | by (metis field_differentiable_within_subset subsetD) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 311 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 312 | lemma holomorphic_transform: "\<lbrakk>f holomorphic_on s; \<And>x. x \<in> s \<Longrightarrow> f x = g x\<rbrakk> \<Longrightarrow> g holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 313 | by (metis field_differentiable_transform_within linordered_field_no_ub holomorphic_on_def) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 314 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 315 | 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 | 316 | by (metis holomorphic_transform) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 317 | |
| 62217 | 318 | lemma holomorphic_on_linear [simp, holomorphic_intros]: "(op * c) holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 319 | unfolding holomorphic_on_def by (metis field_differentiable_linear) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 320 | |
| 62217 | 321 | lemma holomorphic_on_const [simp, holomorphic_intros]: "(\<lambda>z. c) holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 322 | unfolding holomorphic_on_def by (metis field_differentiable_const) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 323 | |
| 62217 | 324 | lemma holomorphic_on_ident [simp, holomorphic_intros]: "(\<lambda>x. x) holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 325 | unfolding holomorphic_on_def by (metis field_differentiable_ident) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 326 | |
| 62217 | 327 | lemma holomorphic_on_id [simp, holomorphic_intros]: "id holomorphic_on s" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 328 | unfolding id_def by (rule holomorphic_on_ident) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 329 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 330 | lemma holomorphic_on_compose: | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 331 | "f holomorphic_on s \<Longrightarrow> g holomorphic_on (f ` s) \<Longrightarrow> (g o f) holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 332 | using field_differentiable_compose_within[of f _ s g] | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 333 | by (auto simp: holomorphic_on_def) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 334 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 335 | lemma holomorphic_on_compose_gen: | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 336 | "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 | 337 | by (metis holomorphic_on_compose holomorphic_on_subset) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 338 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 339 | lemma holomorphic_on_minus [holomorphic_intros]: "f holomorphic_on s \<Longrightarrow> (\<lambda>z. -(f z)) holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 340 | by (metis field_differentiable_minus holomorphic_on_def) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 341 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 342 | lemma holomorphic_on_add [holomorphic_intros]: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 343 | "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z + g z) holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 344 | unfolding holomorphic_on_def by (metis field_differentiable_add) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 345 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 346 | lemma holomorphic_on_diff [holomorphic_intros]: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 347 | "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z - g z) holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 348 | unfolding holomorphic_on_def by (metis field_differentiable_diff) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 349 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 350 | lemma holomorphic_on_mult [holomorphic_intros]: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 351 | "\<lbrakk>f holomorphic_on s; g holomorphic_on s\<rbrakk> \<Longrightarrow> (\<lambda>z. f z * g z) holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 352 | unfolding holomorphic_on_def by (metis field_differentiable_mult) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 353 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 354 | lemma holomorphic_on_inverse [holomorphic_intros]: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 355 | "\<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" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 356 | unfolding holomorphic_on_def by (metis field_differentiable_inverse) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 357 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 358 | lemma holomorphic_on_divide [holomorphic_intros]: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 359 | "\<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" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 360 | unfolding holomorphic_on_def by (metis field_differentiable_divide) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 361 | |
| 61520 
8f85bb443d33
Cauchy's integral formula, required lemmas, and a bit of reorganisation
 paulson <lp15@cam.ac.uk> parents: 
61518diff
changeset | 362 | lemma holomorphic_on_power [holomorphic_intros]: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 363 | "f holomorphic_on s \<Longrightarrow> (\<lambda>z. (f z)^n) holomorphic_on s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 364 | unfolding holomorphic_on_def by (metis field_differentiable_power) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 365 | |
| 64267 | 366 | lemma holomorphic_on_sum [holomorphic_intros]: | 
| 367 | "(\<And>i. i \<in> I \<Longrightarrow> (f i) holomorphic_on s) \<Longrightarrow> (\<lambda>x. sum (\<lambda>i. f i x) I) holomorphic_on s" | |
| 368 | unfolding holomorphic_on_def by (metis field_differentiable_sum) | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 369 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 370 | lemma DERIV_deriv_iff_field_differentiable: | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 371 | "DERIV f x :> deriv f x \<longleftrightarrow> f field_differentiable at x" | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 372 | unfolding field_differentiable_def by (metis DERIV_imp_deriv) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 373 | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 374 | lemma holomorphic_derivI: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 375 | "\<lbrakk>f holomorphic_on S; open S; x \<in> S\<rbrakk> | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 376 | \<Longrightarrow> (f has_field_derivative deriv f x) (at x within T)" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 377 | by (metis DERIV_deriv_iff_field_differentiable at_within_open holomorphic_on_def has_field_derivative_at_within) | 
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 378 | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 379 | lemma complex_derivative_chain: | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 380 | "f field_differentiable at x \<Longrightarrow> g field_differentiable at (f x) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 381 | \<Longrightarrow> deriv (g o f) x = deriv g (f x) * deriv f x" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 382 | by (metis DERIV_deriv_iff_field_differentiable DERIV_chain DERIV_imp_deriv) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 383 | |
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62217diff
changeset | 384 | lemma deriv_linear [simp]: "deriv (\<lambda>w. c * w) = (\<lambda>z. c)" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 385 | by (metis DERIV_imp_deriv DERIV_cmult_Id) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 386 | |
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62217diff
changeset | 387 | lemma deriv_ident [simp]: "deriv (\<lambda>w. w) = (\<lambda>z. 1)" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 388 | by (metis DERIV_imp_deriv DERIV_ident) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 389 | |
| 62397 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62217diff
changeset | 390 | lemma deriv_id [simp]: "deriv id = (\<lambda>z. 1)" | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62217diff
changeset | 391 | by (simp add: id_def) | 
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62217diff
changeset | 392 | |
| 
5ae24f33d343
Substantial new material for multivariate analysis. Also removal of some duplicates.
 paulson <lp15@cam.ac.uk> parents: 
62217diff
changeset | 393 | lemma deriv_const [simp]: "deriv (\<lambda>w. c) = (\<lambda>z. 0)" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 394 | by (metis DERIV_imp_deriv DERIV_const) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 395 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 396 | lemma deriv_add [simp]: | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 397 | "\<lbrakk>f field_differentiable at z; g field_differentiable at z\<rbrakk> | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 398 | \<Longrightarrow> deriv (\<lambda>w. f w + g w) z = deriv f z + deriv g z" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 399 | unfolding DERIV_deriv_iff_field_differentiable[symmetric] | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 400 | by (auto intro!: DERIV_imp_deriv derivative_intros) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 401 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 402 | lemma deriv_diff [simp]: | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 403 | "\<lbrakk>f field_differentiable at z; g field_differentiable at z\<rbrakk> | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 404 | \<Longrightarrow> deriv (\<lambda>w. f w - g w) z = deriv f z - deriv g z" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 405 | unfolding DERIV_deriv_iff_field_differentiable[symmetric] | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 406 | by (auto intro!: DERIV_imp_deriv derivative_intros) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 407 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 408 | lemma deriv_mult [simp]: | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 409 | "\<lbrakk>f field_differentiable at z; g field_differentiable at z\<rbrakk> | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 410 | \<Longrightarrow> deriv (\<lambda>w. f w * g w) z = f z * deriv g z + deriv f z * g z" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 411 | unfolding DERIV_deriv_iff_field_differentiable[symmetric] | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 412 | by (auto intro!: DERIV_imp_deriv derivative_eq_intros) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 413 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 414 | lemma deriv_cmult [simp]: | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 415 | "f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. c * f w) z = c * deriv f z" | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 416 | unfolding DERIV_deriv_iff_field_differentiable[symmetric] | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 417 | by (auto intro!: DERIV_imp_deriv derivative_eq_intros) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 418 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 419 | lemma deriv_cmult_right [simp]: | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 420 | "f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. f w * c) z = deriv f z * c" | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 421 | unfolding DERIV_deriv_iff_field_differentiable[symmetric] | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 422 | by (auto intro!: DERIV_imp_deriv derivative_eq_intros) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 423 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 424 | lemma deriv_cdivide_right [simp]: | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 425 | "f field_differentiable at z \<Longrightarrow> deriv (\<lambda>w. f w / c) z = deriv f z / c" | 
| 62217 | 426 | unfolding Fields.field_class.field_divide_inverse | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 427 | by (blast intro: deriv_cmult_right) | 
| 62217 | 428 | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 429 | lemma complex_derivative_transform_within_open: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 430 | "\<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> | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 431 | \<Longrightarrow> deriv f z = deriv g z" | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 432 | unfolding holomorphic_on_def | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 433 | by (rule DERIV_imp_deriv) | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 434 | (metis DERIV_deriv_iff_field_differentiable DERIV_transform_within_open at_within_open) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 435 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 436 | lemma deriv_compose_linear: | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 437 | "f field_differentiable at (c * z) \<Longrightarrow> deriv (\<lambda>w. f (c * w)) z = c * deriv f (c * z)" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 438 | apply (rule DERIV_imp_deriv) | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 439 | apply (simp add: DERIV_deriv_iff_field_differentiable [symmetric]) | 
| 59554 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
58877diff
changeset | 440 | apply (drule DERIV_chain' [of "times c" c z UNIV f "deriv f (c * z)", OF DERIV_cmult_Id]) | 
| 
4044f53326c9
inlined rules to free user-space from technical names
 haftmann parents: 
58877diff
changeset | 441 | apply (simp add: algebra_simps) | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 442 | done | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 443 | |
| 62533 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 444 | lemma nonzero_deriv_nonconstant: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 445 | assumes df: "DERIV f \<xi> :> df" and S: "open S" "\<xi> \<in> S" and "df \<noteq> 0" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 446 | shows "\<not> f constant_on S" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 447 | unfolding constant_on_def | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 448 | by (metis \<open>df \<noteq> 0\<close> DERIV_transform_within_open [OF df S] DERIV_const DERIV_unique) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 449 | |
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 450 | lemma holomorphic_nonconstant: | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 451 | assumes holf: "f holomorphic_on S" and "open S" "\<xi> \<in> S" "deriv f \<xi> \<noteq> 0" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 452 | shows "\<not> f constant_on S" | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 453 | apply (rule nonzero_deriv_nonconstant [of f "deriv f \<xi>" \<xi> S]) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 454 | using assms | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 455 | apply (auto simp: holomorphic_derivI) | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 456 | done | 
| 
bc25f3916a99
new material to Blochj's theorem, as well as supporting lemmas
 paulson <lp15@cam.ac.uk> parents: 
62408diff
changeset | 457 | |
| 64394 | 458 | subsection\<open>Caratheodory characterization\<close> | 
| 459 | ||
| 460 | lemma field_differentiable_caratheodory_at: | |
| 461 | "f field_differentiable (at z) \<longleftrightarrow> | |
| 462 | (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z) g)" | |
| 463 | using CARAT_DERIV [of f] | |
| 464 | by (simp add: field_differentiable_def has_field_derivative_def) | |
| 465 | ||
| 466 | lemma field_differentiable_caratheodory_within: | |
| 467 | "f field_differentiable (at z within s) \<longleftrightarrow> | |
| 468 | (\<exists>g. (\<forall>w. f(w) - f(z) = g(w) * (w - z)) \<and> continuous (at z within s) g)" | |
| 469 | using DERIV_caratheodory_within [of f] | |
| 470 | by (simp add: field_differentiable_def has_field_derivative_def) | |
| 471 | ||
| 60420 | 472 | subsection\<open>Analyticity on a set\<close> | 
| 56215 | 473 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 474 | definition analytic_on (infixl "(analytic'_on)" 50) | 
| 56215 | 475 | where | 
| 476 | "f analytic_on s \<equiv> \<forall>x \<in> s. \<exists>e. 0 < e \<and> f holomorphic_on (ball x e)" | |
| 477 | ||
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 478 | named_theorems analytic_intros "introduction rules for proving analyticity" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 479 | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 480 | 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 | 481 | by (simp add: at_within_open [OF _ open_ball] analytic_on_def holomorphic_on_def) | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 482 | (metis centre_in_ball field_differentiable_at_within) | 
| 56215 | 483 | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 484 | lemma analytic_on_open: "open s \<Longrightarrow> f analytic_on s \<longleftrightarrow> f holomorphic_on s" | 
| 56215 | 485 | apply (auto simp: analytic_imp_holomorphic) | 
| 486 | apply (auto simp: analytic_on_def holomorphic_on_def) | |
| 487 | by (metis holomorphic_on_def holomorphic_on_subset open_contains_ball) | |
| 488 | ||
| 489 | lemma analytic_on_imp_differentiable_at: | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 490 | "f analytic_on s \<Longrightarrow> x \<in> s \<Longrightarrow> f field_differentiable (at x)" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 491 | apply (auto simp: analytic_on_def holomorphic_on_def) | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 492 | by (metis Topology_Euclidean_Space.open_ball centre_in_ball field_differentiable_within_open) | 
| 56215 | 493 | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 494 | lemma analytic_on_subset: "f analytic_on s \<Longrightarrow> t \<subseteq> s \<Longrightarrow> f analytic_on t" | 
| 56215 | 495 | by (auto simp: analytic_on_def) | 
| 496 | ||
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 497 | lemma analytic_on_Un: "f analytic_on (s \<union> t) \<longleftrightarrow> f analytic_on s \<and> f analytic_on t" | 
| 56215 | 498 | by (auto simp: analytic_on_def) | 
| 499 | ||
| 60585 | 500 | lemma analytic_on_Union: "f analytic_on (\<Union>s) \<longleftrightarrow> (\<forall>t \<in> s. f analytic_on t)" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 501 | by (auto simp: analytic_on_def) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 502 | |
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 503 | lemma analytic_on_UN: "f analytic_on (\<Union>i\<in>I. s i) \<longleftrightarrow> (\<forall>i\<in>I. f analytic_on (s i))" | 
| 56215 | 504 | by (auto simp: analytic_on_def) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 505 | |
| 56215 | 506 | lemma analytic_on_holomorphic: | 
| 507 | "f analytic_on s \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f holomorphic_on t)" | |
| 508 | (is "?lhs = ?rhs") | |
| 509 | proof - | |
| 510 | have "?lhs \<longleftrightarrow> (\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t)" | |
| 511 | proof safe | |
| 512 | assume "f analytic_on s" | |
| 513 | then show "\<exists>t. open t \<and> s \<subseteq> t \<and> f analytic_on t" | |
| 514 | apply (simp add: analytic_on_def) | |
| 515 |       apply (rule exI [where x="\<Union>{u. open u \<and> f analytic_on u}"], auto)
 | |
| 516 | apply (metis Topology_Euclidean_Space.open_ball analytic_on_open centre_in_ball) | |
| 517 | by (metis analytic_on_def) | |
| 518 | next | |
| 519 | fix t | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 520 | assume "open t" "s \<subseteq> t" "f analytic_on t" | 
| 56215 | 521 | then show "f analytic_on s" | 
| 522 | by (metis analytic_on_subset) | |
| 523 | qed | |
| 524 | also have "... \<longleftrightarrow> ?rhs" | |
| 525 | by (auto simp: analytic_on_open) | |
| 526 | finally show ?thesis . | |
| 527 | qed | |
| 528 | ||
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 529 | lemma analytic_on_linear [analytic_intros,simp]: "(op * c) analytic_on s" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 530 | by (auto simp add: analytic_on_holomorphic) | 
| 56215 | 531 | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 532 | lemma analytic_on_const [analytic_intros,simp]: "(\<lambda>z. c) analytic_on s" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 533 | by (metis analytic_on_def holomorphic_on_const zero_less_one) | 
| 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 534 | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 535 | lemma analytic_on_ident [analytic_intros,simp]: "(\<lambda>x. x) analytic_on s" | 
| 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 536 | by (simp add: analytic_on_def gt_ex) | 
| 56215 | 537 | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 538 | lemma analytic_on_id [analytic_intros]: "id analytic_on s" | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 539 | unfolding id_def by (rule analytic_on_ident) | 
| 56215 | 540 | |
| 541 | lemma analytic_on_compose: | |
| 542 | assumes f: "f analytic_on s" | |
| 543 | and g: "g analytic_on (f ` s)" | |
| 544 | shows "(g o f) analytic_on s" | |
| 545 | unfolding analytic_on_def | |
| 546 | proof (intro ballI) | |
| 547 | fix x | |
| 548 | assume x: "x \<in> s" | |
| 549 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball x e" using f | |
| 550 | by (metis analytic_on_def) | |
| 551 | obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball (f x) e'" using g | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 552 | by (metis analytic_on_def g image_eqI x) | 
| 56215 | 553 | have "isCont f x" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 554 | by (metis analytic_on_imp_differentiable_at field_differentiable_imp_continuous_at f x) | 
| 56215 | 555 | with e' obtain d where d: "0 < d" and fd: "f ` ball x d \<subseteq> ball (f x) e'" | 
| 556 | by (auto simp: continuous_at_ball) | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 557 | have "g \<circ> f holomorphic_on ball x (min d e)" | 
| 56215 | 558 | apply (rule holomorphic_on_compose) | 
| 559 | apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 560 | by (metis fd gh holomorphic_on_subset image_mono min.cobounded1 subset_ball) | |
| 561 | then show "\<exists>e>0. g \<circ> f holomorphic_on ball x e" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 562 | by (metis d e min_less_iff_conj) | 
| 56215 | 563 | qed | 
| 564 | ||
| 565 | lemma analytic_on_compose_gen: | |
| 566 | "f analytic_on s \<Longrightarrow> g analytic_on t \<Longrightarrow> (\<And>z. z \<in> s \<Longrightarrow> f z \<in> t) | |
| 567 | \<Longrightarrow> g o f analytic_on s" | |
| 568 | by (metis analytic_on_compose analytic_on_subset image_subset_iff) | |
| 569 | ||
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 570 | lemma analytic_on_neg [analytic_intros]: | 
| 56215 | 571 | "f analytic_on s \<Longrightarrow> (\<lambda>z. -(f z)) analytic_on s" | 
| 572 | by (metis analytic_on_holomorphic holomorphic_on_minus) | |
| 573 | ||
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 574 | lemma analytic_on_add [analytic_intros]: | 
| 56215 | 575 | assumes f: "f analytic_on s" | 
| 576 | and g: "g analytic_on s" | |
| 577 | shows "(\<lambda>z. f z + g z) analytic_on s" | |
| 578 | unfolding analytic_on_def | |
| 579 | proof (intro ballI) | |
| 580 | fix z | |
| 581 | assume z: "z \<in> s" | |
| 582 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f | |
| 583 | by (metis analytic_on_def) | |
| 584 | obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 585 | by (metis analytic_on_def g z) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 586 | have "(\<lambda>z. f z + g z) holomorphic_on ball z (min e e')" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 587 | apply (rule holomorphic_on_add) | 
| 56215 | 588 | apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | 
| 589 | by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 590 | then show "\<exists>e>0. (\<lambda>z. f z + g z) holomorphic_on ball z e" | |
| 591 | by (metis e e' min_less_iff_conj) | |
| 592 | qed | |
| 593 | ||
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 594 | lemma analytic_on_diff [analytic_intros]: | 
| 56215 | 595 | assumes f: "f analytic_on s" | 
| 596 | and g: "g analytic_on s" | |
| 597 | shows "(\<lambda>z. f z - g z) analytic_on s" | |
| 598 | unfolding analytic_on_def | |
| 599 | proof (intro ballI) | |
| 600 | fix z | |
| 601 | assume z: "z \<in> s" | |
| 602 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f | |
| 603 | by (metis analytic_on_def) | |
| 604 | obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 605 | by (metis analytic_on_def g z) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 606 | have "(\<lambda>z. f z - g z) holomorphic_on ball z (min e e')" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 607 | apply (rule holomorphic_on_diff) | 
| 56215 | 608 | apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | 
| 609 | by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 610 | then show "\<exists>e>0. (\<lambda>z. f z - g z) holomorphic_on ball z e" | |
| 611 | by (metis e e' min_less_iff_conj) | |
| 612 | qed | |
| 613 | ||
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 614 | lemma analytic_on_mult [analytic_intros]: | 
| 56215 | 615 | assumes f: "f analytic_on s" | 
| 616 | and g: "g analytic_on s" | |
| 617 | shows "(\<lambda>z. f z * g z) analytic_on s" | |
| 618 | unfolding analytic_on_def | |
| 619 | proof (intro ballI) | |
| 620 | fix z | |
| 621 | assume z: "z \<in> s" | |
| 622 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f | |
| 623 | by (metis analytic_on_def) | |
| 624 | obtain e' where e': "0 < e'" and gh: "g holomorphic_on ball z e'" using g | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 625 | by (metis analytic_on_def g z) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 626 | have "(\<lambda>z. f z * g z) holomorphic_on ball z (min e e')" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 627 | apply (rule holomorphic_on_mult) | 
| 56215 | 628 | apply (metis fh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | 
| 629 | by (metis gh holomorphic_on_subset min.bounded_iff order_refl subset_ball) | |
| 630 | then show "\<exists>e>0. (\<lambda>z. f z * g z) holomorphic_on ball z e" | |
| 631 | by (metis e e' min_less_iff_conj) | |
| 632 | qed | |
| 633 | ||
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 634 | lemma analytic_on_inverse [analytic_intros]: | 
| 56215 | 635 | assumes f: "f analytic_on s" | 
| 636 | and nz: "(\<And>z. z \<in> s \<Longrightarrow> f z \<noteq> 0)" | |
| 637 | shows "(\<lambda>z. inverse (f z)) analytic_on s" | |
| 638 | unfolding analytic_on_def | |
| 639 | proof (intro ballI) | |
| 640 | fix z | |
| 641 | assume z: "z \<in> s" | |
| 642 | then obtain e where e: "0 < e" and fh: "f holomorphic_on ball z e" using f | |
| 643 | by (metis analytic_on_def) | |
| 644 | have "continuous_on (ball z e) f" | |
| 645 | by (metis fh holomorphic_on_imp_continuous_on) | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 646 | then obtain e' where e': "0 < e'" and nz': "\<And>y. dist z y < e' \<Longrightarrow> f y \<noteq> 0" | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 647 | by (metis Topology_Euclidean_Space.open_ball centre_in_ball continuous_on_open_avoid e z nz) | 
| 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 648 | have "(\<lambda>z. inverse (f z)) holomorphic_on ball z (min e e')" | 
| 56215 | 649 | apply (rule holomorphic_on_inverse) | 
| 650 | apply (metis fh holomorphic_on_subset min.cobounded2 min.commute subset_ball) | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 651 | by (metis nz' mem_ball min_less_iff_conj) | 
| 56215 | 652 | then show "\<exists>e>0. (\<lambda>z. inverse (f z)) holomorphic_on ball z e" | 
| 653 | by (metis e e' min_less_iff_conj) | |
| 654 | qed | |
| 655 | ||
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 656 | lemma analytic_on_divide [analytic_intros]: | 
| 56215 | 657 | assumes f: "f analytic_on s" | 
| 658 | and g: "g analytic_on s" | |
| 659 | and nz: "(\<And>z. z \<in> s \<Longrightarrow> g z \<noteq> 0)" | |
| 660 | shows "(\<lambda>z. f z / g z) analytic_on s" | |
| 661 | unfolding divide_inverse | |
| 662 | by (metis analytic_on_inverse analytic_on_mult f g nz) | |
| 663 | ||
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 664 | lemma analytic_on_power [analytic_intros]: | 
| 56215 | 665 | "f analytic_on s \<Longrightarrow> (\<lambda>z. (f z) ^ n) analytic_on s" | 
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 666 | by (induct n) (auto simp: analytic_on_mult) | 
| 56215 | 667 | |
| 65587 
16a8991ab398
New material (and some tidying) purely in the Analysis directory
 paulson <lp15@cam.ac.uk> parents: 
64394diff
changeset | 668 | lemma analytic_on_sum [analytic_intros]: | 
| 64267 | 669 | "(\<And>i. i \<in> I \<Longrightarrow> (f i) analytic_on s) \<Longrightarrow> (\<lambda>x. sum (\<lambda>i. f i x) I) analytic_on s" | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 670 | by (induct I rule: infinite_finite_induct) (auto simp: analytic_on_const analytic_on_add) | 
| 56215 | 671 | |
| 62408 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 672 | lemma deriv_left_inverse: | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 673 | assumes "f holomorphic_on S" and "g holomorphic_on T" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 674 | and "open S" and "open T" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 675 | and "f ` S \<subseteq> T" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 676 | and [simp]: "\<And>z. z \<in> S \<Longrightarrow> g (f z) = z" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 677 | and "w \<in> S" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 678 | shows "deriv f w * deriv g (f w) = 1" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 679 | proof - | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 680 | have "deriv f w * deriv g (f w) = deriv g (f w) * deriv f w" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 681 | by (simp add: algebra_simps) | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 682 | also have "... = deriv (g o f) w" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 683 | using assms | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 684 | by (metis analytic_on_imp_differentiable_at analytic_on_open complex_derivative_chain image_subset_iff) | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 685 | also have "... = deriv id w" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 686 | apply (rule complex_derivative_transform_within_open [where s=S]) | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 687 | apply (rule assms holomorphic_on_compose_gen holomorphic_intros)+ | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 688 | apply simp | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 689 | done | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 690 | also have "... = 1" | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 691 | by simp | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 692 | finally show ?thesis . | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 693 | qed | 
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 694 | |
| 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 695 | subsection\<open>analyticity at a point\<close> | 
| 56215 | 696 | |
| 697 | lemma analytic_at_ball: | |
| 698 |   "f analytic_on {z} \<longleftrightarrow> (\<exists>e. 0<e \<and> f holomorphic_on ball z e)"
 | |
| 699 | by (metis analytic_on_def singleton_iff) | |
| 700 | ||
| 701 | lemma analytic_at: | |
| 702 |     "f analytic_on {z} \<longleftrightarrow> (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s)"
 | |
| 703 | by (metis analytic_on_holomorphic empty_subsetI insert_subset) | |
| 704 | ||
| 705 | lemma analytic_on_analytic_at: | |
| 706 |     "f analytic_on s \<longleftrightarrow> (\<forall>z \<in> s. f analytic_on {z})"
 | |
| 707 | by (metis analytic_at_ball analytic_on_def) | |
| 708 | ||
| 709 | lemma analytic_at_two: | |
| 710 |   "f analytic_on {z} \<and> g analytic_on {z} \<longleftrightarrow>
 | |
| 711 | (\<exists>s. open s \<and> z \<in> s \<and> f holomorphic_on s \<and> g holomorphic_on s)" | |
| 712 | (is "?lhs = ?rhs") | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 713 | proof | 
| 56215 | 714 | assume ?lhs | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 715 | then obtain s t | 
| 56215 | 716 | where st: "open s" "z \<in> s" "f holomorphic_on s" | 
| 717 | "open t" "z \<in> t" "g holomorphic_on t" | |
| 718 | by (auto simp: analytic_at) | |
| 719 | show ?rhs | |
| 720 | apply (rule_tac x="s \<inter> t" in exI) | |
| 721 | using st | |
| 722 | apply (auto simp: Diff_subset holomorphic_on_subset) | |
| 723 | done | |
| 724 | next | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 725 | assume ?rhs | 
| 56215 | 726 | then show ?lhs | 
| 727 | by (force simp add: analytic_at) | |
| 728 | qed | |
| 729 | ||
| 60420 | 730 | subsection\<open>Combining theorems for derivative with ``analytic at'' hypotheses\<close> | 
| 56215 | 731 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 732 | lemma | 
| 56215 | 733 |   assumes "f analytic_on {z}" "g analytic_on {z}"
 | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 734 | 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 | 735 | 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 | 736 | 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 | 737 | f z * deriv g z + deriv f z * g z" | 
| 56215 | 738 | proof - | 
| 739 | obtain s where s: "open s" "z \<in> s" "f holomorphic_on s" "g holomorphic_on s" | |
| 740 | using assms by (metis analytic_at_two) | |
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 741 | 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 | 742 | apply (rule DERIV_imp_deriv [OF DERIV_add]) | 
| 56215 | 743 | using s | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 744 | apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) | 
| 56215 | 745 | done | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 746 | 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 | 747 | apply (rule DERIV_imp_deriv [OF DERIV_diff]) | 
| 56215 | 748 | using s | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 749 | apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) | 
| 56215 | 750 | done | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 751 | 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 | 752 | apply (rule DERIV_imp_deriv [OF DERIV_mult']) | 
| 56215 | 753 | using s | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 754 | apply (auto simp: holomorphic_on_open field_differentiable_def DERIV_deriv_iff_field_differentiable) | 
| 56215 | 755 | done | 
| 756 | qed | |
| 757 | ||
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 758 | lemma deriv_cmult_at: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 759 |   "f analytic_on {z} \<Longrightarrow>  deriv (\<lambda>w. c * f w) z = c * deriv f z"
 | 
| 61848 | 760 | by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const) | 
| 56215 | 761 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 762 | lemma deriv_cmult_right_at: | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 763 |   "f analytic_on {z} \<Longrightarrow>  deriv (\<lambda>w. f w * c) z = deriv f z * c"
 | 
| 61848 | 764 | by (auto simp: complex_derivative_mult_at deriv_const analytic_on_const) | 
| 56215 | 765 | |
| 60420 | 766 | subsection\<open>Complex differentiation of sequences and series\<close> | 
| 56215 | 767 | |
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 768 | (* TODO: Could probably be simplified using Uniform_Limit *) | 
| 56215 | 769 | lemma has_complex_derivative_sequence: | 
| 770 | fixes s :: "complex set" | |
| 771 | assumes cvs: "convex s" | |
| 772 | and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" | |
| 773 | 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" | |
| 61973 | 774 | and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) \<longlongrightarrow> l) sequentially" | 
| 775 | shows "\<exists>g. \<forall>x \<in> s. ((\<lambda>n. f n x) \<longlongrightarrow> g x) sequentially \<and> | |
| 56215 | 776 | (g has_field_derivative (g' x)) (at x within s)" | 
| 777 | proof - | |
| 61973 | 778 | from assms obtain x l where x: "x \<in> s" and tf: "((\<lambda>n. f n x) \<longlongrightarrow> l) sequentially" | 
| 56215 | 779 | by blast | 
| 780 |   { fix e::real assume e: "e > 0"
 | |
| 781 | then obtain N where N: "\<forall>n\<ge>N. \<forall>x. x \<in> s \<longrightarrow> cmod (f' n x - g' x) \<le> e" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 782 | by (metis conv) | 
| 56215 | 783 | 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" | 
| 784 | proof (rule exI [of _ N], clarify) | |
| 785 | fix n y h | |
| 786 | assume "N \<le> n" "y \<in> s" | |
| 787 | then have "cmod (f' n y - g' y) \<le> e" | |
| 788 | by (metis N) | |
| 789 | then have "cmod h * cmod (f' n y - g' y) \<le> cmod h * e" | |
| 790 | by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) | |
| 791 | then show "cmod (f' n y * h - g' y * h) \<le> e * cmod h" | |
| 792 | by (simp add: norm_mult [symmetric] field_simps) | |
| 793 | qed | |
| 794 | } note ** = this | |
| 795 | show ?thesis | |
| 796 | unfolding has_field_derivative_def | |
| 797 | proof (rule has_derivative_sequence [OF cvs _ _ x]) | |
| 798 | show "\<forall>n. \<forall>x\<in>s. (f n has_derivative (op * (f' n x))) (at x within s)" | |
| 799 | by (metis has_field_derivative_def df) | |
| 61969 | 800 | next show "(\<lambda>n. f n x) \<longlonglongrightarrow> l" | 
| 56215 | 801 | by (rule tf) | 
| 802 | 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" | |
| 803 | by (blast intro: **) | |
| 804 | qed | |
| 805 | qed | |
| 806 | ||
| 807 | lemma has_complex_derivative_series: | |
| 808 | fixes s :: "complex set" | |
| 809 | assumes cvs: "convex s" | |
| 810 | and df: "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x within s)" | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 811 | and conv: "\<And>e. 0 < e \<Longrightarrow> \<exists>N. \<forall>n x. n \<ge> N \<longrightarrow> x \<in> s | 
| 56215 | 812 | \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e" | 
| 813 | and "\<exists>x l. x \<in> s \<and> ((\<lambda>n. f n x) sums l)" | |
| 814 | 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))" | |
| 815 | proof - | |
| 816 | from assms obtain x l where x: "x \<in> s" and sf: "((\<lambda>n. f n x) sums l)" | |
| 817 | by blast | |
| 818 |   { fix e::real assume e: "e > 0"
 | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 819 | then obtain N where N: "\<forall>n x. n \<ge> N \<longrightarrow> x \<in> s | 
| 56215 | 820 | \<longrightarrow> cmod ((\<Sum>i<n. f' i x) - g' x) \<le> e" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 821 | by (metis conv) | 
| 56215 | 822 | 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" | 
| 823 | proof (rule exI [of _ N], clarify) | |
| 824 | fix n y h | |
| 825 | assume "N \<le> n" "y \<in> s" | |
| 826 | then have "cmod ((\<Sum>i<n. f' i y) - g' y) \<le> e" | |
| 827 | by (metis N) | |
| 828 | then have "cmod h * cmod ((\<Sum>i<n. f' i y) - g' y) \<le> cmod h * e" | |
| 829 | by (auto simp: antisym_conv2 mult_le_cancel_left norm_triangle_ineq2) | |
| 830 | then show "cmod ((\<Sum>i<n. h * f' i y) - g' y * h) \<le> e * cmod h" | |
| 64267 | 831 | by (simp add: norm_mult [symmetric] field_simps sum_distrib_left) | 
| 56215 | 832 | qed | 
| 833 | } note ** = this | |
| 834 | show ?thesis | |
| 835 | unfolding has_field_derivative_def | |
| 836 | proof (rule has_derivative_series [OF cvs _ _ x]) | |
| 837 | fix n x | |
| 838 | assume "x \<in> s" | |
| 839 | then show "((f n) has_derivative (\<lambda>z. z * f' n x)) (at x within s)" | |
| 840 | by (metis df has_field_derivative_def mult_commute_abs) | |
| 841 | next show " ((\<lambda>n. f n x) sums l)" | |
| 842 | by (rule sf) | |
| 843 | 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" | |
| 844 | by (blast intro: **) | |
| 845 | qed | |
| 846 | qed | |
| 847 | ||
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 848 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 849 | lemma field_differentiable_series: | 
| 66252 | 850 |   fixes f :: "nat \<Rightarrow> 'a::{real_normed_field,banach} \<Rightarrow> 'a"
 | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 851 | assumes "convex s" "open s" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 852 | assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 853 | assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 854 | assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" and x: "x \<in> s" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 855 | shows "summable (\<lambda>n. f n x)" and "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x)" | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 856 | proof - | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 857 | from assms(4) obtain g' where A: "uniform_limit s (\<lambda>n x. \<Sum>i<n. f' i x) g' sequentially" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 858 | unfolding uniformly_convergent_on_def by blast | 
| 61808 | 859 | from x and \<open>open s\<close> have s: "at x within s = at x" by (rule at_within_open) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 860 | have "\<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)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 861 | by (intro has_field_derivative_series[of s f f' g' x0] assms A has_field_derivative_at_within) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 862 | then obtain g where g: "\<And>x. x \<in> s \<Longrightarrow> (\<lambda>n. f n x) sums g x" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 863 | "\<And>x. x \<in> s \<Longrightarrow> (g has_field_derivative g' x) (at x within s)" by blast | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 864 | from g[OF x] show "summable (\<lambda>n. f n x)" by (auto simp: summable_def) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 865 | from g(2)[OF x] have g': "(g has_derivative op * (g' x)) (at x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 866 | by (simp add: has_field_derivative_def s) | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 867 | have "((\<lambda>x. \<Sum>n. f n x) has_derivative op * (g' x)) (at x)" | 
| 62087 
44841d07ef1d
revisions to limits and derivatives, plus new lemmas
 paulson parents: 
61975diff
changeset | 868 | by (rule has_derivative_transform_within_open[OF g' \<open>open s\<close> x]) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 869 | (insert g, auto simp: sums_iff) | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 870 | thus "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x)" unfolding differentiable_def | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 871 | by (auto simp: summable_def field_differentiable_def has_field_derivative_def) | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 872 | qed | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 873 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 874 | lemma field_differentiable_series': | 
| 66252 | 875 |   fixes f :: "nat \<Rightarrow> 'a::{real_normed_field,banach} \<Rightarrow> 'a"
 | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 876 | assumes "convex s" "open s" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 877 | assumes "\<And>n x. x \<in> s \<Longrightarrow> (f n has_field_derivative f' n x) (at x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 878 | assumes "uniformly_convergent_on s (\<lambda>n x. \<Sum>i<n. f' i x)" | 
| 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 879 | assumes "x0 \<in> s" "summable (\<lambda>n. f n x0)" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 880 | shows "(\<lambda>x. \<Sum>n. f n x) field_differentiable (at x0)" | 
| 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 881 | using field_differentiable_series[OF assms, of x0] \<open>x0 \<in> s\<close> by blast+ | 
| 61531 
ab2e862263e7
Rounding function, uniform limits, cotangent, binomial identities
 eberlm parents: 
61520diff
changeset | 882 | |
| 60420 | 883 | subsection\<open>Bound theorem\<close> | 
| 56215 | 884 | |
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 885 | lemma field_differentiable_bound: | 
| 66252 | 886 | fixes s :: "'a::real_normed_field set" | 
| 56215 | 887 | assumes cvs: "convex s" | 
| 888 | and df: "\<And>z. z \<in> s \<Longrightarrow> (f has_field_derivative f' z) (at z within s)" | |
| 889 | and dn: "\<And>z. z \<in> s \<Longrightarrow> norm (f' z) \<le> B" | |
| 890 | and "x \<in> s" "y \<in> s" | |
| 891 | shows "norm(f x - f y) \<le> B * norm(x - y)" | |
| 892 | apply (rule differentiable_bound [OF cvs]) | |
| 56223 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 893 | 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 | 894 | 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 | 895 | apply fact | 
| 
7696903b9e61
generalize theory of operator norms to work with class real_normed_vector
 huffman parents: 
56217diff
changeset | 896 | apply fact | 
| 56215 | 897 | done | 
| 898 | ||
| 62408 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 899 | subsection\<open>Inverse function theorem for complex derivatives\<close> | 
| 56215 | 900 | |
| 66252 | 901 | lemma has_field_derivative_inverse_basic: | 
| 56215 | 902 | shows "DERIV f (g y) :> f' \<Longrightarrow> | 
| 903 | f' \<noteq> 0 \<Longrightarrow> | |
| 904 | continuous (at y) g \<Longrightarrow> | |
| 905 | open t \<Longrightarrow> | |
| 906 | y \<in> t \<Longrightarrow> | |
| 907 | (\<And>z. z \<in> t \<Longrightarrow> f (g z) = z) | |
| 908 | \<Longrightarrow> DERIV g y :> inverse (f')" | |
| 909 | unfolding has_field_derivative_def | |
| 910 | apply (rule has_derivative_inverse_basic) | |
| 911 | apply (auto simp: bounded_linear_mult_right) | |
| 912 | done | |
| 913 | ||
| 66252 | 914 | lemmas has_complex_derivative_inverse_basic = has_field_derivative_inverse_basic | 
| 915 | ||
| 916 | lemma has_field_derivative_inverse_strong: | |
| 917 |   fixes f :: "'a::{euclidean_space,real_normed_field} \<Rightarrow> 'a"
 | |
| 56215 | 918 | shows "DERIV f x :> f' \<Longrightarrow> | 
| 919 | f' \<noteq> 0 \<Longrightarrow> | |
| 920 | open s \<Longrightarrow> | |
| 921 | x \<in> s \<Longrightarrow> | |
| 922 | continuous_on s f \<Longrightarrow> | |
| 923 | (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z) | |
| 924 | \<Longrightarrow> DERIV g (f x) :> inverse (f')" | |
| 925 | unfolding has_field_derivative_def | |
| 926 | apply (rule has_derivative_inverse_strong [of s x f g ]) | |
| 927 | by auto | |
| 66252 | 928 | lemmas has_complex_derivative_inverse_strong = has_field_derivative_inverse_strong | 
| 56215 | 929 | |
| 66252 | 930 | lemma has_field_derivative_inverse_strong_x: | 
| 931 |   fixes f :: "'a::{euclidean_space,real_normed_field} \<Rightarrow> 'a"
 | |
| 56215 | 932 | shows "DERIV f (g y) :> f' \<Longrightarrow> | 
| 933 | f' \<noteq> 0 \<Longrightarrow> | |
| 934 | open s \<Longrightarrow> | |
| 935 | continuous_on s f \<Longrightarrow> | |
| 936 | g y \<in> s \<Longrightarrow> f(g y) = y \<Longrightarrow> | |
| 937 | (\<And>z. z \<in> s \<Longrightarrow> g (f z) = z) | |
| 938 | \<Longrightarrow> DERIV g y :> inverse (f')" | |
| 939 | unfolding has_field_derivative_def | |
| 940 | apply (rule has_derivative_inverse_strong_x [of s g y f]) | |
| 941 | by auto | |
| 66252 | 942 | lemmas has_complex_derivative_inverse_strong_x = has_field_derivative_inverse_strong_x | 
| 56215 | 943 | |
| 60420 | 944 | subsection \<open>Taylor on Complex Numbers\<close> | 
| 56215 | 945 | |
| 64267 | 946 | lemma sum_Suc_reindex: | 
| 56215 | 947 | fixes f :: "nat \<Rightarrow> 'a::ab_group_add" | 
| 64267 | 948 |     shows  "sum f {0..n} = f 0 - f (Suc n) + sum (\<lambda>i. f (Suc i)) {0..n}"
 | 
| 56215 | 949 | by (induct n) auto | 
| 950 | ||
| 66252 | 951 | lemma field_taylor: | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 952 | assumes s: "convex s" | 
| 56215 | 953 | 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)" | 
| 66252 | 954 | and B: "\<And>x. x \<in> s \<Longrightarrow> norm (f (Suc n) x) \<le> B" | 
| 56215 | 955 | and w: "w \<in> s" | 
| 956 | and z: "z \<in> s" | |
| 66252 | 957 | shows "norm(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / (fact i))) | 
| 958 | \<le> B * norm(z - w)^(Suc n) / fact n" | |
| 56215 | 959 | proof - | 
| 960 | have wzs: "closed_segment w z \<subseteq> s" using assms | |
| 961 | by (metis convex_contains_segment) | |
| 962 |   { fix u
 | |
| 963 | assume "u \<in> closed_segment w z" | |
| 964 | then have "u \<in> s" | |
| 965 | by (metis wzs subsetD) | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 966 | have "(\<Sum>i\<le>n. f i u * (- of_nat i * (z-u)^(i - 1)) / (fact i) + | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 967 | f (Suc i) u * (z-u)^i / (fact i)) = | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 968 | f (Suc n) u * (z-u) ^ n / (fact n)" | 
| 56215 | 969 | proof (induction n) | 
| 970 | case 0 show ?case by simp | |
| 971 | next | |
| 972 | case (Suc n) | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 973 | have "(\<Sum>i\<le>Suc n. f i u * (- of_nat i * (z-u) ^ (i - 1)) / (fact i) + | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 974 | f (Suc i) u * (z-u) ^ i / (fact i)) = | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 975 | f (Suc n) u * (z-u) ^ n / (fact n) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 976 | f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n) / (fact (Suc n)) - | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 977 | f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n) / (fact (Suc n))" | 
| 56479 
91958d4b30f7
revert c1bbd3e22226, a14831ac3023, and 36489d77c484: divide_minus_left/right are again simp rules
 hoelzl parents: 
56409diff
changeset | 978 | using Suc by simp | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 979 | also have "... = f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n))" | 
| 56215 | 980 | proof - | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 981 | have "(fact(Suc n)) * | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 982 | (f(Suc n) u *(z-u) ^ n / (fact n) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 983 | f(Suc(Suc n)) u *((z-u) *(z-u) ^ n) / (fact(Suc n)) - | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 984 | f(Suc n) u *((1 + of_nat n) *(z-u) ^ n) / (fact(Suc n))) = | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 985 | ((fact(Suc n)) *(f(Suc n) u *(z-u) ^ n)) / (fact n) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 986 | ((fact(Suc n)) *(f(Suc(Suc n)) u *((z-u) *(z-u) ^ n)) / (fact(Suc n))) - | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 987 | ((fact(Suc n)) *(f(Suc n) u *(of_nat(Suc n) *(z-u) ^ n))) / (fact(Suc n))" | 
| 63367 
6c731c8b7f03
simplified definitions of combinatorial functions
 haftmann parents: 
63332diff
changeset | 988 | by (simp add: algebra_simps del: fact_Suc) | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 989 | also have "... = ((fact (Suc n)) * (f (Suc n) u * (z-u) ^ n)) / (fact n) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 990 | (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 991 | (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" | 
| 63367 
6c731c8b7f03
simplified definitions of combinatorial functions
 haftmann parents: 
63332diff
changeset | 992 | by (simp del: fact_Suc) | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 993 | also have "... = (of_nat (Suc n) * (f (Suc n) u * (z-u) ^ n)) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 994 | (f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)) - | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 995 | (f (Suc n) u * ((1 + of_nat n) * (z-u) ^ n))" | 
| 63367 
6c731c8b7f03
simplified definitions of combinatorial functions
 haftmann parents: 
63332diff
changeset | 996 | by (simp only: fact_Suc of_nat_mult ac_simps) simp | 
| 56215 | 997 | also have "... = f (Suc (Suc n)) u * ((z-u) * (z-u) ^ n)" | 
| 998 | by (simp add: algebra_simps) | |
| 999 | finally show ?thesis | |
| 63367 
6c731c8b7f03
simplified definitions of combinatorial functions
 haftmann parents: 
63332diff
changeset | 1000 | by (simp add: mult_left_cancel [where c = "(fact (Suc n))", THEN iffD1] del: fact_Suc) | 
| 56215 | 1001 | qed | 
| 1002 | finally show ?case . | |
| 1003 | qed | |
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1004 | then have "((\<lambda>v. (\<Sum>i\<le>n. f i v * (z - v)^i / (fact i))) | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1005 | has_field_derivative f (Suc n) u * (z-u) ^ n / (fact n)) | 
| 56215 | 1006 | (at u within s)" | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1007 | apply (intro derivative_eq_intros) | 
| 60420 | 1008 | apply (blast intro: assms \<open>u \<in> s\<close>) | 
| 56215 | 1009 | apply (rule refl)+ | 
| 1010 | apply (auto simp: field_simps) | |
| 1011 | done | |
| 1012 | } note sum_deriv = this | |
| 1013 |   { fix u
 | |
| 1014 | assume u: "u \<in> closed_segment w z" | |
| 1015 | then have us: "u \<in> s" | |
| 1016 | by (metis wzs subsetD) | |
| 66252 | 1017 | have "norm (f (Suc n) u) * norm (z - u) ^ n \<le> norm (f (Suc n) u) * norm (u - z) ^ n" | 
| 56215 | 1018 | by (metis norm_minus_commute order_refl) | 
| 66252 | 1019 | also have "... \<le> norm (f (Suc n) u) * norm (z - w) ^ n" | 
| 56215 | 1020 | by (metis mult_left_mono norm_ge_zero power_mono segment_bound [OF u]) | 
| 66252 | 1021 | also have "... \<le> B * norm (z - w) ^ n" | 
| 56215 | 1022 | by (metis norm_ge_zero zero_le_power mult_right_mono B [OF us]) | 
| 66252 | 1023 | finally have "norm (f (Suc n) u) * norm (z - u) ^ n \<le> B * norm (z - w) ^ n" . | 
| 56215 | 1024 | } note cmod_bound = this | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1025 | have "(\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)) = (\<Sum>i\<le>n. (f i z / (fact i)) * 0 ^ i)" | 
| 56215 | 1026 | by simp | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1027 | also have "\<dots> = f 0 z / (fact 0)" | 
| 64267 | 1028 | by (subst sum_zero_power) simp | 
| 66252 | 1029 | finally have "norm (f 0 z - (\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i))) | 
| 1030 | \<le> norm ((\<Sum>i\<le>n. f i w * (z - w) ^ i / (fact i)) - | |
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1031 | (\<Sum>i\<le>n. f i z * (z - z) ^ i / (fact i)))" | 
| 56215 | 1032 | by (simp add: norm_minus_commute) | 
| 66252 | 1033 | also have "... \<le> B * norm (z - w) ^ n / (fact n) * norm (w - z)" | 
| 62534 
6855b348e828
complex_differentiable -> field_differentiable, etc. (making these theorems also available for type real)
 paulson <lp15@cam.ac.uk> parents: 
62533diff
changeset | 1034 | apply (rule field_differentiable_bound | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1035 | [where f' = "\<lambda>w. f (Suc n) w * (z - w)^n / (fact n)" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61235diff
changeset | 1036 | and s = "closed_segment w z", OF convex_closed_segment]) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1037 | apply (auto simp: ends_in_segment DERIV_subset [OF sum_deriv wzs] | 
| 56215 | 1038 | norm_divide norm_mult norm_power divide_le_cancel cmod_bound) | 
| 1039 | done | |
| 66252 | 1040 | also have "... \<le> B * norm (z - w) ^ Suc n / (fact n)" | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1041 | by (simp add: algebra_simps norm_minus_commute) | 
| 56215 | 1042 | finally show ?thesis . | 
| 1043 | qed | |
| 1044 | ||
| 66252 | 1045 | lemma complex_taylor: | 
| 1046 | assumes s: "convex s" | |
| 1047 | 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)" | |
| 1048 | and B: "\<And>x. x \<in> s \<Longrightarrow> cmod (f (Suc n) x) \<le> B" | |
| 1049 | and w: "w \<in> s" | |
| 1050 | and z: "z \<in> s" | |
| 1051 | shows "cmod(f 0 z - (\<Sum>i\<le>n. f i w * (z-w) ^ i / (fact i))) | |
| 1052 | \<le> B * cmod(z - w)^(Suc n) / fact n" | |
| 1053 | using assms by (rule field_taylor) | |
| 1054 | ||
| 1055 | ||
| 62408 
86f27b264d3d
Conformal_mappings: a big development in complex analysis (+ some lemmas)
 paulson <lp15@cam.ac.uk> parents: 
62397diff
changeset | 1056 | text\<open>Something more like the traditional MVT for real components\<close> | 
| 56370 
7c717ba55a0b
reorder Complex_Analysis_Basics; rename DD to deriv
 hoelzl parents: 
56369diff
changeset | 1057 | |
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1058 | lemma complex_mvt_line: | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1059 | assumes "\<And>u. u \<in> closed_segment w z \<Longrightarrow> (f has_field_derivative f'(u)) (at u)" | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61235diff
changeset | 1060 | shows "\<exists>u. u \<in> closed_segment w z \<and> Re(f z) - Re(f w) = Re(f'(u) * (z - w))" | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1061 | proof - | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1062 | 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 | 1063 | by (simp add: real_vector.scale_left_diff_distrib real_vector.scale_right_diff_distrib) | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1064 | note assms[unfolded has_field_derivative_def, derivative_intros] | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1065 | show ?thesis | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1066 | apply (cut_tac mvt_simple | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1067 | [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 | 1068 | "\<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 | 1069 | apply auto | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1070 | apply (rule_tac x="(1 - x) *\<^sub>R w + x *\<^sub>R z" in exI) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61235diff
changeset | 1071 | apply (auto simp: closed_segment_def twz) [] | 
| 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61235diff
changeset | 1072 | apply (intro derivative_eq_intros has_derivative_at_within, simp_all) | 
| 56369 
2704ca85be98
moved generic theorems from Complex_Analysis_Basic; fixed some theorem names
 hoelzl parents: 
56332diff
changeset | 1073 | apply (simp add: fun_eq_iff real_vector.scale_right_diff_distrib) | 
| 61518 
ff12606337e9
new lemmas about topology, etc., for Cauchy integral formula
 paulson parents: 
61235diff
changeset | 1074 | apply (force simp: twz closed_segment_def) | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1075 | done | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1076 | qed | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1077 | |
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1078 | lemma complex_taylor_mvt: | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1079 | 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 | 1080 | 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 | 1081 | Re (f 0 z) = | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1082 | Re ((\<Sum>i = 0..n. f i w * (z - w) ^ i / (fact i)) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1083 | (f (Suc n) u * (z-u)^n / (fact n)) * (z - w))" | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1084 | proof - | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1085 |   { fix u
 | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1086 | 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 | 1087 | have "(\<Sum>i = 0..n. | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1088 | (f (Suc i) u * (z-u) ^ i - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1089 | (fact i)) = | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1090 | f (Suc 0) u - | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1091 | (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1092 | (fact (Suc n)) + | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1093 | (\<Sum>i = 0..n. | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1094 | (f (Suc (Suc i)) u * ((z-u) ^ Suc i) - of_nat (Suc i) * (f (Suc i) u * (z-u) ^ i)) / | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1095 | (fact (Suc i)))" | 
| 64267 | 1096 | by (subst sum_Suc_reindex) simp | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1097 | also have "... = f (Suc 0) u - | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1098 | (f (Suc (Suc n)) u * ((z-u) ^ Suc n) - (of_nat (Suc n)) * (z-u) ^ n * f (Suc n) u) / | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1099 | (fact (Suc n)) + | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1100 | (\<Sum>i = 0..n. | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1101 | f (Suc (Suc i)) u * ((z-u) ^ Suc i) / (fact (Suc i)) - | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1102 | f (Suc i) u * (z-u) ^ i / (fact i))" | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
56889diff
changeset | 1103 | by (simp only: diff_divide_distrib fact_cancel ac_simps) | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1104 | also have "... = f (Suc 0) u - | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1105 | (f (Suc (Suc n)) u * (z-u) ^ Suc n - of_nat (Suc n) * (z-u) ^ n * f (Suc n) u) / | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1106 | (fact (Suc n)) + | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1107 | f (Suc (Suc n)) u * (z-u) ^ Suc n / (fact (Suc n)) - f (Suc 0) u" | 
| 64267 | 1108 | by (subst sum_Suc_diff) auto | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1109 | also have "... = f (Suc n) u * (z-u) ^ n / (fact n)" | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1110 | by (simp only: algebra_simps diff_divide_distrib fact_cancel) | 
| 61609 
77b453bd616f
Coercion "real" now has type nat => real only and is no longer overloaded. Type class "real_of" is gone. Many duplicate theorems removed.
 paulson <lp15@cam.ac.uk> parents: 
61531diff
changeset | 1111 | finally have "(\<Sum>i = 0..n. (f (Suc i) u * (z - u) ^ i | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1112 | - of_nat i * (f i u * (z-u) ^ (i - Suc 0))) / (fact i)) = | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1113 | f (Suc n) u * (z - u) ^ n / (fact n)" . | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1114 | then have "((\<lambda>u. \<Sum>i = 0..n. f i u * (z - u) ^ i / (fact i)) has_field_derivative | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1115 | f (Suc n) u * (z - u) ^ n / (fact n)) (at u)" | 
| 56381 
0556204bc230
merged DERIV_intros, has_derivative_intros into derivative_intros
 hoelzl parents: 
56371diff
changeset | 1116 | apply (intro derivative_eq_intros)+ | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1117 | apply (force intro: u assms) | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1118 | apply (rule refl)+ | 
| 57514 
bdc2c6b40bf2
prefer ac_simps collections over separate name bindings for add and mult
 haftmann parents: 
56889diff
changeset | 1119 | apply (auto simp: ac_simps) | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1120 | done | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1121 | } | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1122 | then show ?thesis | 
| 59730 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1123 | apply (cut_tac complex_mvt_line [of w z "\<lambda>u. \<Sum>i = 0..n. f i u * (z-u) ^ i / (fact i)" | 
| 
b7c394c7a619
The factorial function, "fact", now has type "nat => 'a"
 paulson <lp15@cam.ac.uk> parents: 
59615diff
changeset | 1124 | "\<lambda>u. (f (Suc n) u * (z-u)^n / (fact n))"]) | 
| 56238 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1125 | 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 | 1126 | done | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1127 | qed | 
| 
5d147e1e18d1
a few new lemmas and generalisations of old ones
 paulson <lp15@cam.ac.uk> parents: 
56223diff
changeset | 1128 | |
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1129 | |
| 60420 | 1130 | subsection \<open>Polynomal function extremal theorem, from HOL Light\<close> | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1131 | |
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1132 | lemma polyfun_extremal_lemma: (*COMPLEX_POLYFUN_EXTREMAL_LEMMA in HOL Light*) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1133 | fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1134 | assumes "0 < e" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1135 | shows "\<exists>M. \<forall>z. M \<le> norm(z) \<longrightarrow> norm (\<Sum>i\<le>n. c(i) * z^i) \<le> e * norm(z) ^ (Suc n)" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1136 | proof (induct n) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1137 | case 0 with assms | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1138 | show ?case | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1139 | apply (rule_tac x="norm (c 0) / e" in exI) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1140 | apply (auto simp: field_simps) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1141 | done | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1142 | next | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1143 | case (Suc n) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1144 | obtain M where M: "\<And>z. M \<le> norm z \<Longrightarrow> norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1145 | using Suc assms by blast | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1146 | show ?case | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1147 | proof (rule exI [where x= "max M (1 + norm(c(Suc n)) / e)"], clarsimp simp del: power_Suc) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1148 | fix z::'a | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1149 | assume z1: "M \<le> norm z" and "1 + norm (c (Suc n)) / e \<le> norm z" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1150 | then have z2: "e + norm (c (Suc n)) \<le> e * norm z" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1151 | using assms by (simp add: field_simps) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1152 | have "norm (\<Sum>i\<le>n. c i * z^i) \<le> e * norm z ^ Suc n" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1153 | using M [OF z1] by simp | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1154 | then have "norm (\<Sum>i\<le>n. c i * z^i) + norm (c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1155 | by simp | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1156 | then have "norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc n + norm (c (Suc n) * z ^ Suc n)" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1157 | by (blast intro: norm_triangle_le elim: ) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1158 | also have "... \<le> (e + norm (c (Suc n))) * norm z ^ Suc n" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1159 | by (simp add: norm_power norm_mult algebra_simps) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1160 | also have "... \<le> (e * norm z) * norm z ^ Suc n" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1161 | by (metis z2 mult.commute mult_left_mono norm_ge_zero norm_power) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1162 | finally show "norm ((\<Sum>i\<le>n. c i * z^i) + c (Suc n) * z ^ Suc n) \<le> e * norm z ^ Suc (Suc n)" | 
| 60162 | 1163 | by simp | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1164 | qed | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1165 | qed | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1166 | |
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1167 | lemma polyfun_extremal: (*COMPLEX_POLYFUN_EXTREMAL in HOL Light*) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1168 | fixes c :: "nat \<Rightarrow> 'a::real_normed_div_algebra" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1169 | assumes k: "c k \<noteq> 0" "1\<le>k" and kn: "k\<le>n" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1170 | shows "eventually (\<lambda>z. norm (\<Sum>i\<le>n. c(i) * z^i) \<ge> B) at_infinity" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1171 | using kn | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1172 | proof (induction n) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1173 | case 0 | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1174 | then show ?case | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1175 | using k by simp | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1176 | next | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1177 | case (Suc m) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1178 | let ?even = ?case | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1179 | show ?even | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1180 | proof (cases "c (Suc m) = 0") | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1181 | case True | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1182 | then show ?even using Suc k | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1183 | by auto (metis antisym_conv less_eq_Suc_le not_le) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1184 | next | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1185 | case False | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1186 | then obtain M where M: | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1187 | "\<And>z. M \<le> norm z \<Longrightarrow> norm (\<Sum>i\<le>m. c i * z^i) \<le> norm (c (Suc m)) / 2 * norm z ^ Suc m" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1188 | using polyfun_extremal_lemma [of "norm(c (Suc m)) / 2" c m] Suc | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1189 | by auto | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1190 | have "\<exists>b. \<forall>z. b \<le> norm z \<longrightarrow> B \<le> norm (\<Sum>i\<le>Suc m. c i * z^i)" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1191 | proof (rule exI [where x="max M (max 1 (\<bar>B\<bar> / (norm(c (Suc m)) / 2)))"], clarsimp simp del: power_Suc) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1192 | fix z::'a | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1193 | assume z1: "M \<le> norm z" "1 \<le> norm z" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1194 | and "\<bar>B\<bar> * 2 / norm (c (Suc m)) \<le> norm z" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1195 | then have z2: "\<bar>B\<bar> \<le> norm (c (Suc m)) * norm z / 2" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1196 | using False by (simp add: field_simps) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1197 | have nz: "norm z \<le> norm z ^ Suc m" | 
| 60420 | 1198 | by (metis \<open>1 \<le> norm z\<close> One_nat_def less_eq_Suc_le power_increasing power_one_right zero_less_Suc) | 
| 60017 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1199 | have *: "\<And>y x. norm (c (Suc m)) * norm z / 2 \<le> norm y - norm x \<Longrightarrow> B \<le> norm (x + y)" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1200 | by (metis abs_le_iff add.commute norm_diff_ineq order_trans z2) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1201 | have "norm z * norm (c (Suc m)) + 2 * norm (\<Sum>i\<le>m. c i * z^i) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1202 | \<le> norm (c (Suc m)) * norm z + norm (c (Suc m)) * norm z ^ Suc m" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1203 | using M [of z] Suc z1 by auto | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1204 | also have "... \<le> 2 * (norm (c (Suc m)) * norm z ^ Suc m)" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1205 | using nz by (simp add: mult_mono del: power_Suc) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1206 | finally show "B \<le> norm ((\<Sum>i\<le>m. c i * z^i) + c (Suc m) * z ^ Suc m)" | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1207 | using Suc.IH | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1208 | apply (auto simp: eventually_at_infinity) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1209 | apply (rule *) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1210 | apply (simp add: field_simps norm_mult norm_power) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1211 | done | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1212 | qed | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1213 | then show ?even | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1214 | by (simp add: eventually_at_infinity) | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1215 | qed | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1216 | qed | 
| 
b785d6d06430
Overloading of ln and powr, but "approximation" no longer works for powr. Code generation also fails due to type ambiguity in scala.
 paulson <lp15@cam.ac.uk> parents: 
59730diff
changeset | 1217 | |
| 56215 | 1218 | end |