src/HOL/Complex_Analysis/Complex_Singularities.thy
author wenzelm
Wed, 12 Mar 2025 11:39:00 +0100
changeset 82265 4b875a4c83b0
parent 81899 1171ea4a23e4
child 82310 41f5266e5595
permissions -rw-r--r--
update for release;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     1
theory Complex_Singularities
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     2
  imports Conformal_Mappings
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     3
begin
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     4
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     5
subsection \<open>Non-essential singular points\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     6
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
     7
definition\<^marker>\<open>tag important\<close>
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
     8
  is_pole :: "('a::topological_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool" 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
     9
  where "is_pole f a =  (LIM x (at a). f x :> at_infinity)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    10
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    11
lemma is_pole_cong:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    12
  assumes "eventually (\<lambda>x. f x = g x) (at a)" "a=b"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    13
  shows "is_pole f a \<longleftrightarrow> is_pole g b"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    14
  unfolding is_pole_def using assms by (intro filterlim_cong,auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    15
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    16
lemma is_pole_transform:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    17
  assumes "is_pole f a" "eventually (\<lambda>x. f x = g x) (at a)" "a=b"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    18
  shows "is_pole g b"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    19
  using is_pole_cong assms by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    20
73795
8893e0ed263a new lemmas mostly about paths
paulson <lp15@cam.ac.uk>
parents: 72222
diff changeset
    21
lemma is_pole_shift_iff:
8893e0ed263a new lemmas mostly about paths
paulson <lp15@cam.ac.uk>
parents: 72222
diff changeset
    22
  fixes f :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector)"
8893e0ed263a new lemmas mostly about paths
paulson <lp15@cam.ac.uk>
parents: 72222
diff changeset
    23
  shows "is_pole (f \<circ> (+) d) a = is_pole f (a + d)"
8893e0ed263a new lemmas mostly about paths
paulson <lp15@cam.ac.uk>
parents: 72222
diff changeset
    24
  by (metis add_diff_cancel_right' filterlim_shift_iff is_pole_def)
8893e0ed263a new lemmas mostly about paths
paulson <lp15@cam.ac.uk>
parents: 72222
diff changeset
    25
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    26
lemma is_pole_tendsto:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
    27
  fixes f:: "('a::topological_space \<Rightarrow> 'b::real_normed_div_algebra)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    28
  shows "is_pole f x \<Longrightarrow> ((inverse o f) \<longlongrightarrow> 0) (at x)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    29
  unfolding is_pole_def
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
    30
  by (auto simp add: filterlim_inverse_at_iff[symmetric] comp_def filterlim_at)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    31
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    32
lemma is_pole_shift_0:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    33
  fixes f :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    34
  shows "is_pole f z \<longleftrightarrow> is_pole (\<lambda>x. f (z + x)) 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    35
  unfolding is_pole_def by (subst at_to_0) (auto simp: filterlim_filtermap add_ac)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    36
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    37
lemma is_pole_shift_0':
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    38
  fixes f :: "('a::real_normed_vector \<Rightarrow> 'b::real_normed_vector)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    39
  shows "NO_MATCH 0 z \<Longrightarrow> is_pole f z \<longleftrightarrow> is_pole (\<lambda>x. f (z + x)) 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    40
  by (metis is_pole_shift_0)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    41
77228
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
    42
lemma is_pole_compose_iff:
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
    43
  assumes "filtermap g (at x) = (at y)"
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
    44
  shows   "is_pole (f \<circ> g) x \<longleftrightarrow> is_pole f y"
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
    45
  unfolding is_pole_def filterlim_def filtermap_compose assms ..
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
    46
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    47
lemma is_pole_inverse_holomorphic:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    48
  assumes "open s"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
    49
    and f_holo: "f holomorphic_on (s-{z})"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
    50
    and pole: "is_pole f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
    51
    and non_z: "\<forall>x\<in>s-{z}. f x\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    52
  shows "(\<lambda>x. if x=z then 0 else inverse (f x)) holomorphic_on s"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    53
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    54
  define g where "g \<equiv> \<lambda>x. if x=z then 0 else inverse (f x)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    55
  have "isCont g z" unfolding isCont_def  using is_pole_tendsto[OF pole]
72222
01397b6e5eb0 small quantifier fixes
paulson <lp15@cam.ac.uk>
parents: 71201
diff changeset
    56
    by (simp add: g_def cong: LIM_cong)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    57
  moreover have "continuous_on (s-{z}) f" using f_holo holomorphic_on_imp_continuous_on by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    58
  hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
    59
    by (auto elim!:continuous_on_inverse simp add: non_z)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    60
  hence "continuous_on (s-{z}) g" unfolding g_def
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
    61
    using continuous_on_eq by fastforce
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    62
  ultimately have "continuous_on s g" using open_delete[OF \<open>open s\<close>] \<open>open s\<close>
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
    63
    by (auto simp add: continuous_on_eq_continuous_at)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    64
  moreover have "(inverse o f) holomorphic_on (s-{z})"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    65
    unfolding comp_def using f_holo
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
    66
    by (auto elim!:holomorphic_on_inverse simp add: non_z)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    67
  hence "g holomorphic_on (s-{z})"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
    68
    using g_def holomorphic_transform by force
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    69
  ultimately show ?thesis unfolding g_def using \<open>open s\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    70
    by (auto elim!: no_isolated_singularity)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    71
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    72
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    73
lemma not_is_pole_holomorphic:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    74
  assumes "open A" "x \<in> A" "f holomorphic_on A"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    75
  shows   "\<not>is_pole f x"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    76
proof -
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    77
  have "continuous_on A f" 
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    78
    by (intro holomorphic_on_imp_continuous_on) fact
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    79
  with assms have "isCont f x" 
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    80
    by (simp add: continuous_on_eq_continuous_at)
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    81
  hence "f \<midarrow>x\<rightarrow> f x" 
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    82
    by (simp add: isCont_def)
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    83
  thus "\<not>is_pole f x" 
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    84
    unfolding is_pole_def
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    85
    using not_tendsto_and_filterlim_at_infinity[of "at x" f "f x"] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    86
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    87
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    88
lemma is_pole_inverse_power: "n > 0 \<Longrightarrow> is_pole (\<lambda>z::complex. 1 / (z - a) ^ n) a"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    89
  unfolding is_pole_def inverse_eq_divide [symmetric]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    90
  by (intro filterlim_compose[OF filterlim_inverse_at_infinity] tendsto_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    91
     (auto simp: filterlim_at eventually_at intro!: exI[of _ 1] tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    92
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    93
lemma is_pole_cmult_iff [simp]:
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    94
  assumes "c \<noteq> 0"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    95
  shows "is_pole (\<lambda>z. c * f z :: 'a :: real_normed_field) z \<longleftrightarrow> is_pole f z"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    96
proof
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    97
  assume "is_pole (\<lambda>z. c * f z) z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    98
  with \<open>c\<noteq>0\<close> have "is_pole (\<lambda>z. inverse c * (c * f z)) z" 
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
    99
    unfolding is_pole_def
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   100
    by (force intro: tendsto_mult_filterlim_at_infinity)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   101
  thus "is_pole f z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   102
    using \<open>c\<noteq>0\<close> by (simp add: field_simps)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   103
next
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   104
  assume "is_pole f z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   105
  with \<open>c\<noteq>0\<close> show "is_pole (\<lambda>z. c * f z) z"  
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   106
    by (auto intro!: tendsto_mult_filterlim_at_infinity simp: is_pole_def)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   107
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   108
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   109
lemma is_pole_uminus_iff [simp]: "is_pole (\<lambda>z. -f z :: 'a :: real_normed_field) z \<longleftrightarrow> is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   110
  using is_pole_cmult_iff[of "-1" f] by (simp del: is_pole_cmult_iff)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   111
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   112
lemma is_pole_inverse: "is_pole (\<lambda>z::complex. 1 / (z - a)) a"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   113
  using is_pole_inverse_power[of 1 a] by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   114
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   115
lemma is_pole_divide:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   116
  fixes f :: "'a :: t2_space \<Rightarrow> 'b :: real_normed_field"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   117
  assumes "isCont f z" "filterlim g (at 0) (at z)" "f z \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   118
  shows   "is_pole (\<lambda>z. f z / g z) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   119
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   120
  have "filterlim (\<lambda>z. f z * inverse (g z)) at_infinity (at z)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   121
    using assms filterlim_compose filterlim_inverse_at_infinity isCont_def
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   122
      tendsto_mult_filterlim_at_infinity by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   123
  thus ?thesis by (simp add: field_split_simps is_pole_def)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   124
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   125
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   126
lemma is_pole_basic:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   127
  assumes "f holomorphic_on A" "open A" "z \<in> A" "f z \<noteq> 0" "n > 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   128
  shows   "is_pole (\<lambda>w. f w / (w-z) ^ n) z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   129
proof (rule is_pole_divide)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   130
  have "continuous_on A f" by (rule holomorphic_on_imp_continuous_on) fact
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   131
  with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   132
  have "filterlim (\<lambda>w. (w-z) ^ n) (nhds 0) (at z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   133
    using assms by (auto intro!: tendsto_eq_intros)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   134
  thus "filterlim (\<lambda>w. (w-z) ^ n) (at 0) (at z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   135
    by (intro filterlim_atI tendsto_eq_intros)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   136
       (use assms in \<open>auto simp: eventually_at_filter\<close>)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   137
qed fact+
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   138
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   139
lemma is_pole_basic':
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   140
  assumes "f holomorphic_on A" "open A" "0 \<in> A" "f 0 \<noteq> 0" "n > 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   141
  shows   "is_pole (\<lambda>w. f w / w ^ n) 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   142
  using is_pole_basic[of f A 0] assms by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   143
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   144
lemma is_pole_compose: 
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   145
  assumes "is_pole f w" "g \<midarrow>z\<rightarrow> w" "eventually (\<lambda>z. g z \<noteq> w) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   146
  shows   "is_pole (\<lambda>x. f (g x)) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   147
  using assms(1) unfolding is_pole_def
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   148
  by (rule filterlim_compose) (use assms in \<open>auto simp: filterlim_at\<close>)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   149
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   150
lemma is_pole_plus_const_iff:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   151
  "is_pole f z \<longleftrightarrow> is_pole (\<lambda>x. f x + c) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   152
proof 
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   153
  assume "is_pole f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   154
  then have "filterlim f at_infinity (at z)" unfolding is_pole_def .
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   155
  moreover have "((\<lambda>_. c) \<longlongrightarrow> c) (at z)" by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   156
  ultimately have " LIM x (at z). f x + c :> at_infinity"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   157
    using tendsto_add_filterlim_at_infinity'[of f "at z"] by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   158
  then show "is_pole (\<lambda>x. f x + c) z" unfolding is_pole_def .
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   159
next
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   160
  assume "is_pole (\<lambda>x. f x + c) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   161
  then have "filterlim (\<lambda>x. f x + c) at_infinity (at z)" 
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   162
    unfolding is_pole_def .
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   163
  moreover have "((\<lambda>_. -c) \<longlongrightarrow> -c) (at z)" by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   164
  ultimately have "LIM x (at z). f x :> at_infinity"
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   165
    using tendsto_add_filterlim_at_infinity'[of "(\<lambda>x. f x + c)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   166
        "at z" "(\<lambda>_. - c)" "-c"] 
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   167
    by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   168
  then show "is_pole f z" unfolding is_pole_def .
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   169
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   170
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   171
lemma is_pole_minus_const_iff:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   172
  "is_pole (\<lambda>x. f x - c) z \<longleftrightarrow> is_pole f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   173
  using is_pole_plus_const_iff [of f z "-c"] by simp
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   174
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   175
lemma is_pole_alt:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   176
  "is_pole f x  = (\<forall>B>0. \<exists>U. open U \<and> x\<in>U \<and> (\<forall>y\<in>U. y\<noteq>x \<longrightarrow> norm (f y)\<ge>B))"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   177
  unfolding is_pole_def
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   178
  unfolding filterlim_at_infinity[of 0,simplified] eventually_at_topological
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   179
  by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   180
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   181
lemma is_pole_mult_analytic_nonzero1:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   182
  assumes "is_pole g x" "f analytic_on {x}" "f x \<noteq> 0"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   183
  shows   "is_pole (\<lambda>x. f x * g x) x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   184
  unfolding is_pole_def
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   185
proof (rule tendsto_mult_filterlim_at_infinity)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   186
  show "f \<midarrow>x\<rightarrow> f x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   187
    using assms by (simp add: analytic_at_imp_isCont isContD)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   188
qed (use assms in \<open>auto simp: is_pole_def\<close>)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   189
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   190
lemma is_pole_mult_analytic_nonzero2:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   191
  assumes "is_pole f x" "g analytic_on {x}" "g x \<noteq> 0"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   192
  shows   "is_pole (\<lambda>x. f x * g x) x"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   193
proof -
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   194
  have g: "g analytic_on {x}"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   195
    using assms by auto
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   196
  show ?thesis
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   197
    using is_pole_mult_analytic_nonzero1 [OF \<open>is_pole f x\<close> g] \<open>g x \<noteq> 0\<close>
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   198
    by (simp add: mult.commute)
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   199
qed
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   200
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   201
lemma is_pole_mult_analytic_nonzero1_iff:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   202
  assumes "f analytic_on {x}" "f x \<noteq> 0"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   203
  shows   "is_pole (\<lambda>x. f x * g x) x \<longleftrightarrow> is_pole g x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   204
proof
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   205
  assume "is_pole g x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   206
  thus "is_pole (\<lambda>x. f x * g x) x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   207
    by (intro is_pole_mult_analytic_nonzero1 assms)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   208
next
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   209
  assume "is_pole (\<lambda>x. f x * g x) x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   210
  hence "is_pole (\<lambda>x. inverse (f x) * (f x * g x)) x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   211
    by (rule is_pole_mult_analytic_nonzero1)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   212
       (use assms in \<open>auto intro!: analytic_intros\<close>)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   213
  also have "?this \<longleftrightarrow> is_pole g x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   214
  proof (rule is_pole_cong)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   215
    have "eventually (\<lambda>x. f x \<noteq> 0) (at x)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   216
      using assms by (simp add: analytic_at_neq_imp_eventually_neq)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   217
    thus "eventually (\<lambda>x. inverse (f x) * (f x * g x) = g x) (at x)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   218
      by eventually_elim auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   219
  qed auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   220
  finally show "is_pole g x" .
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   221
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   222
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   223
lemma is_pole_mult_analytic_nonzero2_iff:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   224
  assumes "g analytic_on {x}" "g x \<noteq> 0"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   225
  shows   "is_pole (\<lambda>x. f x * g x) x \<longleftrightarrow> is_pole f x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   226
  by (subst mult.commute, rule is_pole_mult_analytic_nonzero1_iff) (fact assms)+
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   227
77228
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   228
lemma frequently_const_imp_not_is_pole:
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   229
  fixes z :: "'a::first_countable_topology"
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   230
  assumes "frequently (\<lambda>w. f w = c) (at z)"
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   231
  shows   "\<not> is_pole f z"
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   232
proof
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   233
  assume "is_pole f z"
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   234
  from assms have "z islimpt {w. f w = c}"
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   235
    by (simp add: islimpt_conv_frequently_at)
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   236
  then obtain g where g: "\<And>n. g n \<in> {w. f w = c} - {z}" "g \<longlonglongrightarrow> z"
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   237
    unfolding islimpt_sequential by blast
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   238
  then have "(f \<circ> g) \<longlonglongrightarrow> c"
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   239
    by (simp add: tendsto_eventually)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   240
  moreover have "filterlim g (at z) sequentially"
77228
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   241
    using g by (auto simp: filterlim_at)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   242
  then have "filterlim (f \<circ> g) at_infinity sequentially"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   243
    unfolding o_def
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   244
    using \<open>is_pole f z\<close> filterlim_compose is_pole_def by blast
77228
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   245
  ultimately show False
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   246
    using not_tendsto_and_filterlim_at_infinity trivial_limit_sequentially by blast
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   247
qed
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   248
  
8c093a4b8ccf Even more new material from Eberl and Li
paulson <lp15@cam.ac.uk>
parents: 77226
diff changeset
   249
 text \<open>The proposition
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   250
              \<^term>\<open>\<exists>x. ((f::complex\<Rightarrow>complex) \<longlongrightarrow> x) (at z) \<or> is_pole f z\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   251
can be interpreted as the complex function \<^term>\<open>f\<close> has a non-essential singularity at \<^term>\<open>z\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   252
(i.e. the singularity is either removable or a pole).\<close>
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   253
definition not_essential:: "[complex \<Rightarrow> complex, complex] \<Rightarrow> bool" where
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   254
  "not_essential f z = (\<exists>x. f\<midarrow>z\<rightarrow>x \<or> is_pole f z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   255
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   256
definition isolated_singularity_at:: "[complex \<Rightarrow> complex, complex] \<Rightarrow> bool" where
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   257
  "isolated_singularity_at f z = (\<exists>r>0. f analytic_on ball z r-{z})"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   258
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   259
lemma not_essential_cong:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   260
  assumes "eventually (\<lambda>x. f x = g x) (at z)" "z = z'"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   261
  shows   "not_essential f z \<longleftrightarrow> not_essential g z'"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   262
  unfolding not_essential_def using assms filterlim_cong is_pole_cong by fastforce
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   263
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents: 77228
diff changeset
   264
lemma not_essential_compose_iff:
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents: 77228
diff changeset
   265
  assumes "filtermap g (at z) = at z'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents: 77228
diff changeset
   266
  shows   "not_essential (f \<circ> g) z = not_essential f z'"
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents: 77228
diff changeset
   267
  unfolding not_essential_def filterlim_def filtermap_compose assms is_pole_compose_iff[OF assms]
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents: 77228
diff changeset
   268
  by blast
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents: 77228
diff changeset
   269
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   270
lemma isolated_singularity_at_cong:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   271
  assumes "eventually (\<lambda>x. f x = g x) (at z)" "z = z'"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   272
  shows   "isolated_singularity_at f z \<longleftrightarrow> isolated_singularity_at g z'"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   273
proof -
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   274
  have "isolated_singularity_at g z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   275
    if "isolated_singularity_at f z" "eventually (\<lambda>x. f x = g x) (at z)" for f g
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   276
  proof -
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   277
    from that(1) obtain r where r: "r > 0" "f analytic_on ball z r - {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   278
      by (auto simp: isolated_singularity_at_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   279
    from that(2) obtain r' where r': "r' > 0" "\<forall>x\<in>ball z r'-{z}. f x = g x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   280
      unfolding eventually_at_filter eventually_nhds_metric by (auto simp: dist_commute)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   281
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   282
    have "f holomorphic_on ball z r - {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   283
      using r(2) by (subst (asm) analytic_on_open) auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   284
    hence "f holomorphic_on ball z (min r r') - {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   285
      by (rule holomorphic_on_subset) auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   286
    also have "?this \<longleftrightarrow> g holomorphic_on ball z (min r r') - {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   287
      using r' by (intro holomorphic_cong) auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   288
    also have "\<dots> \<longleftrightarrow> g analytic_on ball z (min r r') - {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   289
      by (subst analytic_on_open) auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   290
    finally show ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   291
      unfolding isolated_singularity_at_def
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   292
      by (intro exI[of _ "min r r'"]) (use \<open>r > 0\<close> \<open>r' > 0\<close> in auto)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   293
  qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   294
  from this[of f g] this[of g f] assms show ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   295
    by (auto simp: eq_commute)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   296
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   297
  
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   298
lemma removable_singularity:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   299
  assumes "f holomorphic_on A - {x}" "open A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   300
  assumes "f \<midarrow>x\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   301
  shows   "(\<lambda>y. if y = x then c else f y) holomorphic_on A" (is "?g holomorphic_on _")
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   302
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   303
  have "continuous_on A ?g"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   304
    unfolding continuous_on_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   305
  proof
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   306
    fix y assume y: "y \<in> A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   307
    show "(?g \<longlongrightarrow> ?g y) (at y within A)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   308
    proof (cases "y = x")
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   309
      case False
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   310
      have "continuous_on (A - {x}) f"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   311
        using assms(1) by (meson holomorphic_on_imp_continuous_on)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   312
      hence "(f \<longlongrightarrow> ?g y) (at y within A - {x})"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   313
        using False y by (auto simp: continuous_on_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   314
      also have "?this \<longleftrightarrow> (?g \<longlongrightarrow> ?g y) (at y within A - {x})"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   315
        by (intro filterlim_cong refl) (auto simp: eventually_at_filter)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   316
      also have "at y within A - {x} = at y within A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   317
        using y assms False
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   318
        by (intro at_within_nhd[of _ "A - {x}"]) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   319
      finally show ?thesis .
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   320
    next
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   321
      case [simp]: True
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   322
      have "f \<midarrow>x\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   323
        by fact
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   324
      also have "?this \<longleftrightarrow> (?g \<longlongrightarrow> ?g y) (at y)"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   325
        by (simp add: LIM_equal)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   326
      finally show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   327
        by (meson Lim_at_imp_Lim_at_within)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   328
    qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   329
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   330
  moreover {
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   331
    have "?g holomorphic_on A - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   332
      using assms(1) holomorphic_transform by fastforce
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   333
  }
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   334
  ultimately show ?thesis
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   335
    using assms by (simp add: no_isolated_singularity)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   336
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   337
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   338
lemma removable_singularity':
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   339
  assumes "isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   340
  assumes "f \<midarrow>z\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   341
  shows   "(\<lambda>y. if y = z then c else f y) analytic_on {z}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   342
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   343
  from assms obtain r where r: "r > 0" "f analytic_on ball z r - {z}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   344
    by (auto simp: isolated_singularity_at_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   345
  have "(\<lambda>y. if y = z then c else f y) holomorphic_on ball z r"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   346
  proof (rule removable_singularity)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   347
    show "f holomorphic_on ball z r - {z}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   348
      using r(2) by (subst (asm) analytic_on_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   349
  qed (use assms in auto)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   350
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   351
    using r(1) unfolding analytic_at by (intro exI[of _ "ball z r"]) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   352
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   353
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   354
named_theorems singularity_intros "introduction rules for singularities"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   355
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   356
lemma holomorphic_factor_unique:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   357
  fixes f:: "complex \<Rightarrow> complex" and z::complex and r::real and m n::int
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   358
  assumes "r>0" "g z\<noteq>0" "h z\<noteq>0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   359
    and asm: "\<forall>w\<in>ball z r-{z}. f w = g w * (w-z) powi n \<and> g w\<noteq>0 \<and> f w =  h w * (w-z) powi m \<and> h w\<noteq>0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   360
    and g_holo: "g holomorphic_on ball z r" and h_holo: "h holomorphic_on ball z r"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   361
  shows "n=m"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   362
proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   363
  have [simp]: "at z within ball z r \<noteq> bot" using \<open>r>0\<close>
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   364
      by (auto simp add: at_within_ball_bot_iff)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   365
  have False when "n>m"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   366
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   367
    have "(h \<longlongrightarrow> 0) (at z within ball z r)"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   368
    proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w-z) powi (n - m) * g w"])
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   369
      have "\<forall>w\<in>ball z r-{z}. h w = (w-z)powi(n-m) * g w"
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   370
        using \<open>n>m\<close> asm \<open>r>0\<close> by (simp add: field_simps power_int_diff) force
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   371
      then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   372
            \<Longrightarrow> (x' - z) powi (n - m) * g x' = h x'" for x' by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   373
    next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   374
      define F where "F \<equiv> at z within ball z r"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   375
      define f' where "f' \<equiv> \<lambda>x. (x - z) powi (n-m)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   376
      have "f' z=0" using \<open>n>m\<close> unfolding f'_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   377
      moreover have "continuous F f'" unfolding f'_def F_def continuous_def
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   378
        using \<open>n>m\<close>
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   379
          by (auto simp add: Lim_ident_at  intro!:tendsto_powr_complex_0 tendsto_eq_intros)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   380
      ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   381
        by (simp add: continuous_within)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   382
      moreover have "(g \<longlongrightarrow> g z) F"
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents: 77228
diff changeset
   383
        unfolding F_def
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents: 77228
diff changeset
   384
        using \<open>r>0\<close> centre_in_ball continuous_on_def g_holo holomorphic_on_imp_continuous_on by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   385
      ultimately show " ((\<lambda>w. f' w * g w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   386
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   387
    moreover have "(h \<longlongrightarrow> h z) (at z within ball z r)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   388
      using holomorphic_on_imp_continuous_on[OF h_holo]
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   389
      by (auto simp add: continuous_on_def \<open>r>0\<close>)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   390
    ultimately have "h z=0" by (auto intro!: tendsto_unique)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   391
    thus False using \<open>h z\<noteq>0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   392
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   393
  moreover have False when "m>n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   394
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   395
    have "(g \<longlongrightarrow> 0) (at z within ball z r)"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   396
    proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w-z) powi (m - n) * h w"])
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   397
      have "\<forall>w\<in>ball z r -{z}. g w = (w-z) powi (m-n) * h w" using \<open>m>n\<close> asm
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   398
        by (simp add: field_simps power_int_diff) force
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   399
      then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   400
            \<Longrightarrow> (x' - z) powi (m - n) * h x' = g x'" for x' by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   401
    next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   402
      define F where "F \<equiv> at z within ball z r"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   403
      define f' where "f' \<equiv>\<lambda>x. (x - z) powi (m-n)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   404
      have "f' z=0" using \<open>m>n\<close> unfolding f'_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   405
      moreover have "continuous F f'" unfolding f'_def F_def continuous_def
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   406
        using \<open>m>n\<close>
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   407
        by (auto simp: Lim_ident_at intro!:tendsto_powr_complex_0 tendsto_eq_intros)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   408
      ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   409
        by (simp add: continuous_within)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   410
      moreover have "(h \<longlongrightarrow> h z) F"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   411
        using holomorphic_on_imp_continuous_on[OF h_holo,unfolded continuous_on_def] \<open>r>0\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   412
        unfolding F_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   413
      ultimately show " ((\<lambda>w. f' w * h w) \<longlongrightarrow> 0) F" using tendsto_mult by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   414
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   415
    moreover have "(g \<longlongrightarrow> g z) (at z within ball z r)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   416
      using holomorphic_on_imp_continuous_on[OF g_holo]
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   417
      by (auto simp add: continuous_on_def \<open>r>0\<close>)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   418
    ultimately have "g z=0" by (auto intro!: tendsto_unique)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   419
    thus False using \<open>g z\<noteq>0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   420
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   421
  ultimately show "n=m" by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   422
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   423
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   424
lemma holomorphic_factor_puncture:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   425
  assumes f_iso: "isolated_singularity_at f z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   426
      and "not_essential f z" \<comment> \<open>\<^term>\<open>f\<close> has either a removable singularity or a pole at \<^term>\<open>z\<close>\<close>
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   427
      and non_zero: "\<exists>\<^sub>Fw in (at z). f w\<noteq>0" \<comment> \<open>\<^term>\<open>f\<close> will not be constantly zero in a neighbour of \<^term>\<open>z\<close>\<close>
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   428
  shows "\<exists>!n::int. \<exists>g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   429
          \<and> (\<forall>w\<in>cball z r-{z}. f w = g w * (w-z) powi n \<and> g w\<noteq>0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   430
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   431
  define P where "P = (\<lambda>f n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   432
          \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powi n  \<and> g w\<noteq>0))"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   433
  have imp_unique: "\<exists>!n::int. \<exists>g r. P f n g r" when "\<exists>n g r. P f n g r"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   434
  proof (rule ex_ex1I[OF that])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   435
    fix n1 n2 :: int
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   436
    assume g1_asm: "\<exists>g1 r1. P f n1 g1 r1" and g2_asm: "\<exists>g2 r2. P f n2 g2 r2"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   437
    define fac where "fac \<equiv> \<lambda>n g r. \<forall>w\<in>cball z r-{z}. f w = g w * (w-z) powi n \<and> g w \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   438
    obtain g1 r1 where "0 < r1" and g1_holo: "g1 holomorphic_on cball z r1" and "g1 z\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   439
        and "fac n1 g1 r1" using g1_asm unfolding P_def fac_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   440
    obtain g2 r2 where "0 < r2" and g2_holo: "g2 holomorphic_on cball z r2" and "g2 z\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   441
        and "fac n2 g2 r2" using g2_asm unfolding P_def fac_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   442
    define r where "r \<equiv> min r1 r2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   443
    have "r>0" using \<open>r1>0\<close> \<open>r2>0\<close> unfolding r_def by auto
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   444
    moreover have "\<forall>w\<in>ball z r-{z}. f w = g1 w * (w-z) powi n1 \<and> g1 w\<noteq>0
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   445
        \<and> f w = g2 w * (w-z) powi n2  \<and> g2 w\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   446
      using \<open>fac n1 g1 r1\<close> \<open>fac n2 g2 r2\<close>   unfolding fac_def r_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   447
      by fastforce
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   448
    ultimately show "n1=n2" 
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   449
      using g1_holo g2_holo \<open>g1 z\<noteq>0\<close> \<open>g2 z\<noteq>0\<close>
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   450
      apply (elim holomorphic_factor_unique)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   451
      by (auto simp add: r_def)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   452
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   453
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   454
  have P_exist: "\<exists> n g r. P h n g r" when
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   455
      "\<exists>z'. (h \<longlongrightarrow> z') (at z)" "isolated_singularity_at h z"  "\<exists>\<^sub>Fw in (at z). h w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   456
    for h
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   457
  proof -
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   458
    from that(2) obtain r where "r>0" and r: "h analytic_on ball z r - {z}"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   459
      unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   460
    obtain z' where "(h \<longlongrightarrow> z') (at z)" using \<open>\<exists>z'. (h \<longlongrightarrow> z') (at z)\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   461
    define h' where "h'=(\<lambda>x. if x=z then z' else h x)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   462
    have "h' holomorphic_on ball z r"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   463
    proof (rule no_isolated_singularity'[of "{z}"])
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   464
      show "\<And>w. w \<in> {z} \<Longrightarrow> (h' \<longlongrightarrow> h' w) (at w within ball z r)"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   465
        by (simp add: LIM_cong Lim_at_imp_Lim_at_within \<open>h \<midarrow>z\<rightarrow> z'\<close> h'_def)
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   466
      show "h' holomorphic_on ball z r - {z}"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   467
        using r analytic_imp_holomorphic h'_def holomorphic_transform by fastforce
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   468
    qed auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   469
    have ?thesis when "z'=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   470
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   471
      have "h' z=0" using that unfolding h'_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   472
      moreover have "\<not> h' constant_on ball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   473
        using \<open>\<exists>\<^sub>Fw in (at z). h w\<noteq>0\<close> unfolding constant_on_def frequently_def eventually_at h'_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   474
        by (metis \<open>0 < r\<close> centre_in_ball dist_commute mem_ball that)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   475
      moreover note \<open>h' holomorphic_on ball z r\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   476
      ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \<subseteq> ball z r" and
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   477
          g: "g holomorphic_on ball z r1"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   478
          "\<And>w. w \<in> ball z r1 \<Longrightarrow> h' w = (w-z) ^ n * g w"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   479
          "\<And>w. w \<in> ball z r1 \<Longrightarrow> g w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   480
        using holomorphic_factor_zero_nonconstant[of _ "ball z r" z thesis,simplified,
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   481
                OF \<open>h' holomorphic_on ball z r\<close> \<open>r>0\<close> \<open>h' z=0\<close> \<open>\<not> h' constant_on ball z r\<close>]
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   482
        by (auto simp add: dist_commute)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   483
      define rr where "rr=r1/2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   484
      have "P h' n g rr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   485
        unfolding P_def rr_def
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   486
        using \<open>n>0\<close> \<open>r1>0\<close> g by (auto simp add: powr_nat)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   487
      then have "P h n g rr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   488
        unfolding h'_def P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   489
      then show ?thesis unfolding P_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   490
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   491
    moreover have ?thesis when "z'\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   492
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   493
      have "h' z\<noteq>0" using that unfolding h'_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   494
      obtain r1 where "r1>0" "cball z r1 \<subseteq> ball z r" "\<forall>x\<in>cball z r1. h' x\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   495
      proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   496
        have "isCont h' z" "h' z\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   497
          by (auto simp add: Lim_cong_within \<open>h \<midarrow>z\<rightarrow> z'\<close> \<open>z'\<noteq>0\<close> continuous_at h'_def)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   498
        then obtain r2 where r2: "r2>0" "\<forall>x\<in>ball z r2. h' x\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   499
          using continuous_at_avoid[of z h' 0 ] unfolding ball_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   500
        define r1 where "r1=min r2 r / 2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   501
        have "0 < r1" "cball z r1 \<subseteq> ball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   502
          using \<open>r2>0\<close> \<open>r>0\<close> unfolding r1_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   503
        moreover have "\<forall>x\<in>cball z r1. h' x \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   504
          using r2 unfolding r1_def by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   505
        ultimately show ?thesis using that by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   506
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   507
      then have "P h' 0 h' r1" using \<open>h' holomorphic_on ball z r\<close> unfolding P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   508
      then have "P h 0 h' r1" unfolding P_def h'_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   509
      then show ?thesis unfolding P_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   510
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   511
    ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   512
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   513
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   514
  have ?thesis when "\<exists>x. (f \<longlongrightarrow> x) (at z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   515
    apply (rule_tac imp_unique[unfolded P_def])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   516
    using P_exist[OF that(1) f_iso non_zero] unfolding P_def .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   517
  moreover have ?thesis when "is_pole f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   518
  proof (rule imp_unique[unfolded P_def])
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   519
    obtain e where [simp]: "e>0" and e_holo: "f holomorphic_on ball z e - {z}" and e_nz: "\<forall>x\<in>ball z e-{z}. f x\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   520
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   521
      have "\<forall>\<^sub>F z in at z. f z \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   522
        using \<open>is_pole f z\<close> filterlim_at_infinity_imp_eventually_ne unfolding is_pole_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   523
        by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   524
      then obtain e1 where e1: "e1>0" "\<forall>x\<in>ball z e1-{z}. f x\<noteq>0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   525
        using that eventually_at[of "\<lambda>x. f x\<noteq>0" z UNIV,simplified] by (auto simp add: dist_commute)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   526
      obtain e2 where e2: "e2>0" "f holomorphic_on ball z e2 - {z}"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   527
        using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   528
      show ?thesis
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   529
        using e1 e2 by (force intro: that[of "min e1 e2"])
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   530
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   531
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   532
    define h where "h \<equiv> \<lambda>x. inverse (f x)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   533
    have "\<exists>n g r. P h n g r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   534
    proof -
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   535
      have "(\<lambda>x. inverse (f x)) analytic_on ball z e - {z}"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   536
        by (metis e_holo e_nz open_ball analytic_on_open holomorphic_on_inverse open_delete)
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   537
      moreover have "h \<midarrow>z\<rightarrow> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   538
        using Lim_transform_within_open assms(2) h_def is_pole_tendsto that by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   539
      moreover have "\<exists>\<^sub>Fw in (at z). h w\<noteq>0"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   540
        using non_zero by (simp add: h_def)
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   541
      ultimately show ?thesis
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   542
        using P_exist[of h] \<open>e > 0\<close>
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   543
        unfolding isolated_singularity_at_def h_def
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   544
        by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   545
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   546
    then obtain n g r
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   547
      where "0 < r" and
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   548
            g_holo: "g holomorphic_on cball z r" and "g z\<noteq>0" and
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   549
            g_fac: "(\<forall>w\<in>cball z r-{z}. h w = g w * (w-z) powi n  \<and> g w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   550
      unfolding P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   551
    have "P f (-n) (inverse o g) r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   552
    proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   553
      have "f w = inverse (g w) * (w-z) powi (- n)" when "w\<in>cball z r - {z}" for w
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   554
        by (metis g_fac h_def inverse_inverse_eq inverse_mult_distrib power_int_minus that)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   555
      then show ?thesis
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   556
        unfolding P_def comp_def
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   557
        using \<open>r>0\<close> g_holo g_fac \<open>g z\<noteq>0\<close> by (auto intro: holomorphic_intros)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   558
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   559
    then show "\<exists>x g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z \<noteq> 0
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   560
                  \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powi x  \<and> g w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   561
      unfolding P_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   562
  qed
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   563
  ultimately show ?thesis 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   564
    using \<open>not_essential f z\<close> unfolding not_essential_def by presburger
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   565
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   566
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   567
lemma not_essential_transform:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   568
  assumes "not_essential g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   569
  assumes "\<forall>\<^sub>F w in (at z). g w = f w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   570
  shows "not_essential f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   571
  using assms unfolding not_essential_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   572
  by (simp add: filterlim_cong is_pole_cong)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   573
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   574
lemma isolated_singularity_at_transform:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   575
  assumes "isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   576
  assumes "\<forall>\<^sub>F w in (at z). g w = f w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   577
  shows "isolated_singularity_at f z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   578
  using assms isolated_singularity_at_cong by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   579
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   580
lemma not_essential_powr[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   581
  assumes "LIM w (at z). f w :> (at x)"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   582
  shows "not_essential (\<lambda>w. (f w) powi n) z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   583
proof -
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   584
  define fp where "fp=(\<lambda>w. (f w) powi n)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   585
  have ?thesis when "n>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   586
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   587
    have "(\<lambda>w.  (f w) ^ (nat n)) \<midarrow>z\<rightarrow> x ^ nat n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   588
      using that assms unfolding filterlim_at by (auto intro!:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   589
    then have "fp \<midarrow>z\<rightarrow> x ^ nat n" unfolding fp_def
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   590
      by (smt (verit) LIM_cong power_int_def that)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   591
    then show ?thesis unfolding not_essential_def fp_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   592
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   593
  moreover have ?thesis when "n=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   594
  proof -
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   595
    have "\<forall>\<^sub>F x in at z. fp x = 1"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   596
      using that filterlim_at_within_not_equal[OF assms] by (auto simp: fp_def)
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   597
    then have "fp \<midarrow>z\<rightarrow> 1"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   598
      by (simp add: tendsto_eventually)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   599
    then show ?thesis unfolding fp_def not_essential_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   600
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   601
  moreover have ?thesis when "n<0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   602
  proof (cases "x=0")
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   603
    case True
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   604
    have "(\<lambda>x. f x ^ nat (- n)) \<midarrow>z\<rightarrow> 0"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   605
      using assms True that unfolding filterlim_at by (auto intro!:tendsto_eq_intros)
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   606
    moreover have "\<forall>\<^sub>F x in at z. f x ^ nat (- n) \<noteq> 0"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   607
      by (smt (verit) True assms eventually_at_topological filterlim_at power_eq_0_iff)
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   608
    ultimately have "LIM w (at z). inverse ((f w) ^ (nat (-n))) :> at_infinity"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   609
      by (metis filterlim_atI filterlim_compose filterlim_inverse_at_infinity)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   610
    then have "LIM w (at z). fp w :> at_infinity"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   611
    proof (elim filterlim_mono_eventually)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   612
      show "\<forall>\<^sub>F x in at z. inverse (f x ^ nat (- n)) = fp x"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   613
        using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   614
        by (smt (verit) eventuallyI power_int_def power_inverse that)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   615
    qed auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   616
    then show ?thesis unfolding fp_def not_essential_def is_pole_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   617
  next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   618
    case False
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   619
    let ?xx= "inverse (x ^ (nat (-n)))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   620
    have "(\<lambda>w. inverse ((f w) ^ (nat (-n)))) \<midarrow>z\<rightarrow>?xx"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   621
      using assms False unfolding filterlim_at by (auto intro!:tendsto_eq_intros)
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   622
    then have "fp \<midarrow>z\<rightarrow> ?xx"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   623
      by (smt (verit, best) LIM_cong fp_def power_int_def power_inverse that)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   624
    then show ?thesis unfolding fp_def not_essential_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   625
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   626
  ultimately show ?thesis by linarith
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   627
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   628
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   629
lemma isolated_singularity_at_powr[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   630
  assumes "isolated_singularity_at f z" "\<forall>\<^sub>F w in (at z). f w\<noteq>0"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   631
  shows "isolated_singularity_at (\<lambda>w. (f w) powi n) z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   632
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   633
  obtain r1 where "r1>0" "f analytic_on ball z r1 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   634
    using assms(1) unfolding isolated_singularity_at_def by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   635
  then have r1: "f holomorphic_on ball z r1 - {z}"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   636
    using analytic_on_open[of "ball z r1-{z}" f] by blast
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   637
  obtain r2 where "r2>0" and r2: "\<forall>w. w \<noteq> z \<and> dist w z < r2 \<longrightarrow> f w \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   638
    using assms(2) unfolding eventually_at by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   639
  define r3 where "r3=min r1 r2"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   640
  have "(\<lambda>w. (f w) powi n) holomorphic_on ball z r3 - {z}"
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   641
    by (intro holomorphic_on_power_int) (use r1 r2 in \<open>auto simp: dist_commute r3_def\<close>)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   642
  moreover have "r3>0" unfolding r3_def using \<open>0 < r1\<close> \<open>0 < r2\<close> by linarith
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   643
  ultimately show ?thesis
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   644
    by (meson open_ball analytic_on_open isolated_singularity_at_def open_delete)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   645
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   646
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   647
lemma non_zero_neighbour:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   648
  assumes f_iso: "isolated_singularity_at f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   649
      and f_ness: "not_essential f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   650
      and f_nconst: "\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   651
    shows "\<forall>\<^sub>F w in (at z). f w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   652
proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   653
  obtain fn fp fr
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   654
    where [simp]: "fp z \<noteq> 0" and "fr > 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   655
      and fr: "fp holomorphic_on cball z fr"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   656
              "\<And>w. w \<in> cball z fr - {z} \<Longrightarrow> f w = fp w * (w-z) powi fn \<and> fp w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   657
    using holomorphic_factor_puncture[OF f_iso f_ness f_nconst] by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   658
  have "f w \<noteq> 0" when " w \<noteq> z" "dist w z < fr" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   659
  proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   660
    have "f w = fp w * (w-z) powi fn" "fp w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   661
      using fr that by (auto simp add: dist_commute)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   662
    moreover have "(w-z) powi fn \<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   663
      unfolding powr_eq_0_iff using \<open>w\<noteq>z\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   664
    ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   665
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   666
  then show ?thesis using \<open>fr>0\<close> unfolding eventually_at by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   667
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   668
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   669
lemma non_zero_neighbour_pole:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   670
  assumes "is_pole f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   671
  shows "\<forall>\<^sub>F w in (at z). f w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   672
  using assms filterlim_at_infinity_imp_eventually_ne[of f "at z" 0]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   673
  unfolding is_pole_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   674
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   675
lemma non_zero_neighbour_alt:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   676
  assumes holo: "f holomorphic_on S"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   677
      and "open S" "connected S" "z \<in> S"  "\<beta> \<in> S" "f \<beta> \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   678
    shows "\<forall>\<^sub>F w in (at z). f w\<noteq>0 \<and> w\<in>S"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   679
proof (cases "f z = 0")
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   680
  case True
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   681
  from isolated_zeros[OF holo \<open>open S\<close> \<open>connected S\<close> \<open>z \<in> S\<close> True \<open>\<beta> \<in> S\<close> \<open>f \<beta> \<noteq> 0\<close>]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   682
  obtain r where "0 < r" "ball z r \<subseteq> S" "\<forall>w \<in> ball z r - {z}.f w \<noteq> 0" by metis
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   683
  then show ?thesis
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   684
    by (smt (verit) open_ball centre_in_ball eventually_at_topological insertE insert_Diff subsetD)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   685
next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   686
  case False
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   687
  obtain r1 where r1: "r1>0" "\<forall>y. dist z y < r1 \<longrightarrow> f y \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   688
    using continuous_at_avoid[of z f, OF _ False] assms continuous_on_eq_continuous_at
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   689
      holo holomorphic_on_imp_continuous_on by blast
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   690
  obtain r2 where r2: "r2>0" "ball z r2 \<subseteq> S"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   691
    using assms openE by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   692
  show ?thesis unfolding eventually_at
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   693
    by (metis (no_types) dist_commute order.strict_trans linorder_less_linear mem_ball r1 r2 subsetD)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   694
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   695
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   696
lemma not_essential_times[singularity_intros]:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   697
  assumes f_ness: "not_essential f z" and g_ness: "not_essential g z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   698
  assumes f_iso: "isolated_singularity_at f z" and g_iso: "isolated_singularity_at g z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   699
  shows "not_essential (\<lambda>w. f w * g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   700
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   701
  define fg where "fg = (\<lambda>w. f w * g w)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   702
  have ?thesis when "\<not> ((\<exists>\<^sub>Fw in (at z). f w\<noteq>0) \<and> (\<exists>\<^sub>Fw in (at z). g w\<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   703
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   704
    have "\<forall>\<^sub>Fw in (at z). fg w=0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   705
      using fg_def frequently_elim1 not_eventually that by fastforce
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   706
    from tendsto_cong[OF this] have "fg \<midarrow>z\<rightarrow>0" by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   707
    then show ?thesis unfolding not_essential_def fg_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   708
  qed
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   709
  moreover have ?thesis when f_nconst: "\<exists>\<^sub>Fw in (at z). f w\<noteq>0" and g_nconst: "\<exists>\<^sub>Fw in (at z). g w\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   710
  proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   711
    obtain fn fp fr where [simp]: "fp z \<noteq> 0" and "fr > 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   712
          and fr: "fp holomorphic_on cball z fr"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   713
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w-z) powi fn \<and> fp w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   714
      using holomorphic_factor_puncture[OF f_iso f_ness f_nconst] by auto
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   715
    obtain gn gp gr where [simp]: "gp z \<noteq> 0" and "gr > 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   716
          and gr: "gp holomorphic_on cball z gr"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   717
                  "\<forall>w\<in>cball z gr - {z}. g w = gp w * (w-z) powi gn \<and> gp w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   718
      using holomorphic_factor_puncture[OF g_iso g_ness g_nconst] by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   719
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   720
    define r1 where "r1=(min fr gr)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   721
    have "r1>0" unfolding r1_def using  \<open>fr>0\<close> \<open>gr>0\<close> by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   722
    have fg_times: "fg w = (fp w * gp w) * (w-z) powi (fn+gn)" and fgp_nz: "fp w*gp w\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   723
      when "w\<in>ball z r1 - {z}" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   724
    proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   725
      have "f w = fp w * (w-z) powi fn" "fp w\<noteq>0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   726
        using fr that unfolding r1_def by auto
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   727
      moreover have "g w = gp w * (w-z) powi gn" "gp w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   728
        using gr that unfolding r1_def by auto
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   729
      ultimately show "fg w = (fp w * gp w) * (w-z) powi (fn+gn)" "fp w*gp w\<noteq>0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   730
        using that by (auto simp add: fg_def power_int_add)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   731
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   732
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   733
    obtain [intro]: "fp \<midarrow>z\<rightarrow>fp z" "gp \<midarrow>z\<rightarrow>gp z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   734
        using fr(1) \<open>fr>0\<close> gr(1) \<open>gr>0\<close>
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   735
        by (metis centre_in_ball continuous_at continuous_on_interior
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   736
            holomorphic_on_imp_continuous_on interior_cball)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   737
    have ?thesis when "fn+gn>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   738
    proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   739
      have "(\<lambda>w. (fp w * gp w) * (w-z) ^ (nat (fn+gn))) \<midarrow>z\<rightarrow>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   740
        using that by (auto intro!:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   741
      then have "fg \<midarrow>z\<rightarrow> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   742
        using Lim_transform_within[OF _ \<open>r1>0\<close>]
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   743
        by (smt (verit, best) Diff_iff dist_commute fg_times mem_ball power_int_def singletonD that zero_less_dist_iff)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   744
      then show ?thesis unfolding not_essential_def fg_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   745
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   746
    moreover have ?thesis when "fn+gn=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   747
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   748
      have "(\<lambda>w. fp w * gp w) \<midarrow>z\<rightarrow>fp z*gp z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   749
        using that by (auto intro!:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   750
      then have "fg \<midarrow>z\<rightarrow> fp z*gp z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   751
        apply (elim Lim_transform_within[OF _ \<open>r1>0\<close>])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   752
        apply (subst fg_times)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   753
        by (auto simp add: dist_commute that)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   754
      then show ?thesis unfolding not_essential_def fg_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   755
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   756
    moreover have ?thesis when "fn+gn<0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   757
    proof -
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   758
      have "LIM x at z. (x - z) ^ nat (- (fn + gn)) :> at 0"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   759
        using eventually_at_topological that
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   760
        by (force intro!: tendsto_eq_intros filterlim_atI)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   761
      moreover have "\<exists>c. (\<lambda>c. fp c * gp c) \<midarrow>z\<rightarrow> c \<and> 0 \<noteq> c"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   762
        using \<open>fp \<midarrow>z\<rightarrow> fp z\<close> \<open>gp \<midarrow>z\<rightarrow> gp z\<close> tendsto_mult by fastforce
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   763
      ultimately have "LIM w (at z). fp w * gp w / (w-z)^nat (-(fn+gn)) :> at_infinity"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   764
        using filterlim_divide_at_infinity by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   765
      then have "is_pole fg z" unfolding is_pole_def
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   766
        apply (elim filterlim_transform_within[OF _ _ \<open>r1>0\<close>])
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   767
        using that
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
   768
        by (simp_all add: dist_commute fg_times of_int_of_nat divide_simps power_int_def del: minus_add_distrib)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   769
      then show ?thesis unfolding not_essential_def fg_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   770
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   771
    ultimately show ?thesis unfolding not_essential_def fg_def by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   772
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   773
  ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   774
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   775
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   776
lemma not_essential_inverse[singularity_intros]:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   777
  assumes f_ness: "not_essential f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   778
  assumes f_iso: "isolated_singularity_at f z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   779
  shows "not_essential (\<lambda>w. inverse (f w)) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   780
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   781
  define vf where "vf = (\<lambda>w. inverse (f w))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   782
  have ?thesis when "\<not>(\<exists>\<^sub>Fw in (at z). f w\<noteq>0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   783
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   784
    have "\<forall>\<^sub>Fw in (at z). f w=0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   785
      using not_eventually that by fastforce
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   786
    then have "vf \<midarrow>z\<rightarrow>0" 
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   787
      unfolding vf_def by (simp add: tendsto_eventually)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   788
    then show ?thesis 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   789
      unfolding not_essential_def vf_def by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   790
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   791
  moreover have ?thesis when "is_pole f z"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   792
    by (metis (mono_tags, lifting) filterlim_at filterlim_inverse_at_iff is_pole_def
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   793
        not_essential_def that)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   794
  moreover have ?thesis when "\<exists>x. f\<midarrow>z\<rightarrow>x " and f_nconst: "\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   795
  proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   796
    from that obtain fz where fz: "f\<midarrow>z\<rightarrow>fz" by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   797
    have ?thesis when "fz=0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   798
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   799
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   800
      have "(\<lambda>w. inverse (vf w)) \<midarrow>z\<rightarrow>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   801
        using fz that unfolding vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   802
      moreover have "\<forall>\<^sub>F w in at z. inverse (vf w) \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   803
        using non_zero_neighbour[OF f_iso f_ness f_nconst]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   804
        unfolding vf_def by auto
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   805
      ultimately show ?thesis unfolding not_essential_def vf_def
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   806
         using filterlim_atI filterlim_inverse_at_iff is_pole_def by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   807
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   808
    moreover have ?thesis when "fz\<noteq>0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   809
      using fz not_essential_def tendsto_inverse that by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   810
    ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   811
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   812
  ultimately show ?thesis using f_ness unfolding not_essential_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   813
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   814
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   815
lemma isolated_singularity_at_inverse[singularity_intros]:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   816
  assumes f_iso: "isolated_singularity_at f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   817
      and f_ness: "not_essential f z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   818
  shows "isolated_singularity_at (\<lambda>w. inverse (f w)) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   819
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   820
  define vf where "vf = (\<lambda>w. inverse (f w))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   821
  have ?thesis when "\<not>(\<exists>\<^sub>Fw in (at z). f w\<noteq>0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   822
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   823
    have "\<forall>\<^sub>Fw in (at z). f w=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   824
      using that[unfolded frequently_def, simplified] by (auto elim: eventually_rev_mp)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   825
    then have "\<forall>\<^sub>Fw in (at z). vf w=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   826
      unfolding vf_def by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   827
    then obtain d1 where "d1>0" and d1: "\<forall>x. x \<noteq> z \<and> dist x z < d1 \<longrightarrow> vf x = 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   828
      unfolding eventually_at by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   829
    then have "vf holomorphic_on ball z d1-{z}"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   830
      using holomorphic_transform[of "\<lambda>_. 0"]
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   831
      by (metis Diff_iff dist_commute holomorphic_on_const insert_iff mem_ball)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   832
    then have "vf analytic_on ball z d1 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   833
      by (simp add: analytic_on_open open_delete)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   834
    then show ?thesis 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   835
      using \<open>d1>0\<close> unfolding isolated_singularity_at_def vf_def by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   836
  qed
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   837
  moreover have ?thesis when f_nconst: "\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   838
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   839
    have "\<forall>\<^sub>F w in at z. f w \<noteq> 0" using non_zero_neighbour[OF f_iso f_ness f_nconst] .
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   840
    then obtain d1 where d1: "d1>0" "\<forall>x. x \<noteq> z \<and> dist x z < d1 \<longrightarrow> f x \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   841
      unfolding eventually_at by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   842
    obtain d2 where "d2>0" and d2: "f analytic_on ball z d2 - {z}"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   843
      using f_iso unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   844
    define d3 where "d3=min d1 d2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   845
    have "d3>0" unfolding d3_def using \<open>d1>0\<close> \<open>d2>0\<close> by auto
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   846
    moreover
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   847
    have "f analytic_on ball z d3 - {z}"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   848
      by (smt (verit, best) Diff_iff analytic_on_analytic_at d2 d3_def mem_ball)
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   849
    then have "vf analytic_on ball z d3 - {z}"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   850
      unfolding vf_def
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   851
      by (intro analytic_on_inverse; simp add: d1(2) d3_def dist_commute)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   852
    ultimately show ?thesis 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   853
      unfolding isolated_singularity_at_def vf_def by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   854
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   855
  ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   856
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   857
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   858
lemma not_essential_divide[singularity_intros]:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   859
  assumes f_ness: "not_essential f z" and g_ness: "not_essential g z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   860
  assumes f_iso: "isolated_singularity_at f z" and g_iso: "isolated_singularity_at g z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   861
  shows "not_essential (\<lambda>w. f w / g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   862
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   863
  have "not_essential (\<lambda>w. f w * inverse (g w)) z"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   864
    by (simp add: f_iso f_ness g_iso g_ness isolated_singularity_at_inverse
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   865
        not_essential_inverse not_essential_times)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   866
  then show ?thesis by (simp add: field_simps)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   867
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   868
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   869
lemma
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   870
  assumes f_iso: "isolated_singularity_at f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   871
      and g_iso: "isolated_singularity_at g z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   872
    shows isolated_singularity_at_times[singularity_intros]:
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   873
              "isolated_singularity_at (\<lambda>w. f w * g w) z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   874
      and isolated_singularity_at_add[singularity_intros]:
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   875
              "isolated_singularity_at (\<lambda>w. f w + g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   876
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   877
  obtain d1 d2 where "d1>0" "d2>0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   878
      and d1: "f analytic_on ball z d1 - {z}" and d2: "g analytic_on ball z d2 - {z}"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   879
    using f_iso g_iso unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   880
  define d3 where "d3=min d1 d2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   881
  have "d3>0" unfolding d3_def using \<open>d1>0\<close> \<open>d2>0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   882
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   883
  have fan: "f analytic_on ball z d3 - {z}"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   884
    by (smt (verit, best) Diff_iff analytic_on_analytic_at d1 d3_def mem_ball)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   885
  have gan: "g analytic_on ball z d3 - {z}"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   886
    by (smt (verit, best) Diff_iff analytic_on_analytic_at d2 d3_def mem_ball)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   887
  have "(\<lambda>w. f w * g w) analytic_on ball z d3 - {z}"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   888
    using analytic_on_mult fan gan by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   889
  then show "isolated_singularity_at (\<lambda>w. f w * g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   890
    using \<open>d3>0\<close> unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   891
  have "(\<lambda>w. f w + g w) analytic_on ball z d3 - {z}"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   892
    using analytic_on_add fan gan by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   893
  then show "isolated_singularity_at (\<lambda>w. f w + g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   894
    using \<open>d3>0\<close> unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   895
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   896
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   897
lemma isolated_singularity_at_uminus[singularity_intros]:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
   898
  assumes f_iso: "isolated_singularity_at f z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   899
  shows "isolated_singularity_at (\<lambda>w. - f w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   900
  using assms unfolding isolated_singularity_at_def using analytic_on_neg by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   901
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   902
lemma isolated_singularity_at_id[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   903
     "isolated_singularity_at (\<lambda>w. w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   904
  unfolding isolated_singularity_at_def by (simp add: gt_ex)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   905
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   906
lemma isolated_singularity_at_minus[singularity_intros]:
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   907
  assumes "isolated_singularity_at f z" and "isolated_singularity_at g z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   908
  shows "isolated_singularity_at (\<lambda>w. f w - g w) z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   909
  unfolding diff_conv_add_uminus
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   910
  using assms isolated_singularity_at_add isolated_singularity_at_uminus by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   911
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   912
lemma isolated_singularity_at_divide[singularity_intros]:
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   913
  assumes "isolated_singularity_at f z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   914
      and "isolated_singularity_at g z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   915
      and "not_essential g z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   916
    shows "isolated_singularity_at (\<lambda>w. f w / g w) z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   917
  unfolding divide_inverse
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
   918
  by (simp add: assms isolated_singularity_at_inverse isolated_singularity_at_times)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   919
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   920
lemma isolated_singularity_at_const[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   921
    "isolated_singularity_at (\<lambda>w. c) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   922
  unfolding isolated_singularity_at_def by (simp add: gt_ex)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   923
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   924
lemma isolated_singularity_at_holomorphic:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   925
  assumes "f holomorphic_on s-{z}" "open s" "z\<in>s"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   926
  shows "isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   927
  using assms unfolding isolated_singularity_at_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   928
  by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   929
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   930
lemma isolated_singularity_at_altdef:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   931
  "isolated_singularity_at f z \<longleftrightarrow> eventually (\<lambda>z. f analytic_on {z}) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   932
proof
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   933
  assume "isolated_singularity_at f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   934
  then obtain r where r: "r > 0" "f analytic_on ball z r - {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   935
    unfolding isolated_singularity_at_def by blast
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   936
  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   937
    using r(1) by (intro eventually_at_in_open) auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   938
  thus "eventually (\<lambda>z. f analytic_on {z}) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   939
    by eventually_elim (use r analytic_on_subset in auto)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   940
next
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   941
  assume "eventually (\<lambda>z. f analytic_on {z}) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   942
  then obtain A where A: "open A" "z \<in> A" "\<And>w. w \<in> A - {z} \<Longrightarrow> f analytic_on {w}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   943
    unfolding eventually_at_topological by blast
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   944
  then show "isolated_singularity_at f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   945
    by (meson analytic_imp_holomorphic analytic_on_analytic_at isolated_singularity_at_holomorphic)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   946
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
   947
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   948
lemma isolated_singularity_at_shift:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   949
  assumes "isolated_singularity_at (\<lambda>x. f (x + w)) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   950
  shows   "isolated_singularity_at f (z + w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   951
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   952
  from assms obtain r where r: "r > 0" and ana: "(\<lambda>x. f (x + w)) analytic_on ball z r - {z}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   953
    unfolding isolated_singularity_at_def by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   954
  have "((\<lambda>x. f (x + w)) \<circ> (\<lambda>x. x - w)) analytic_on (ball (z + w) r - {z + w})"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   955
    by (rule analytic_on_compose_gen[OF _ ana])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   956
       (auto simp: dist_norm algebra_simps intro!: analytic_intros)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   957
  hence "f analytic_on (ball (z + w) r - {z + w})"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   958
    by (simp add: o_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   959
  thus ?thesis using r
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   960
    unfolding isolated_singularity_at_def by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   961
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   962
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   963
lemma isolated_singularity_at_shift_iff:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   964
  "isolated_singularity_at f (z + w) \<longleftrightarrow> isolated_singularity_at (\<lambda>x. f (x + w)) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   965
  using isolated_singularity_at_shift[of f w z]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   966
        isolated_singularity_at_shift[of "\<lambda>x. f (x + w)" "-w" "w + z"]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   967
  by (auto simp: algebra_simps)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   968
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   969
lemma isolated_singularity_at_shift_0:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   970
  "NO_MATCH 0 z \<Longrightarrow> isolated_singularity_at f z \<longleftrightarrow> isolated_singularity_at (\<lambda>x. f (z + x)) 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   971
  using isolated_singularity_at_shift_iff[of f 0 z] by (simp add: add_ac)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   972
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   973
lemma not_essential_shift:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   974
  assumes "not_essential (\<lambda>x. f (x + w)) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   975
  shows   "not_essential f (z + w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   976
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   977
  from assms consider c where "(\<lambda>x. f (x + w)) \<midarrow>z\<rightarrow> c" | "is_pole (\<lambda>x. f (x + w)) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   978
    unfolding not_essential_def by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   979
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   980
  proof cases
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   981
    case (1 c)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   982
    hence "f \<midarrow>z + w\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   983
      by (smt (verit, ccfv_SIG) LIM_cong add.assoc filterlim_at_to_0)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   984
    thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   985
      by (auto simp: not_essential_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   986
  next
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   987
    case 2
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   988
    hence "is_pole f (z + w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   989
      by (subst is_pole_shift_iff [symmetric]) (auto simp: o_def add_ac)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   990
    thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   991
      by (auto simp: not_essential_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   992
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   993
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   994
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   995
lemma not_essential_shift_iff: "not_essential f (z + w) \<longleftrightarrow> not_essential (\<lambda>x. f (x + w)) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   996
  using not_essential_shift[of f w z]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   997
        not_essential_shift[of "\<lambda>x. f (x + w)" "-w" "w + z"]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   998
  by (auto simp: algebra_simps)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   999
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1000
lemma not_essential_shift_0:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1001
  "NO_MATCH 0 z \<Longrightarrow> not_essential f z \<longleftrightarrow> not_essential (\<lambda>x. f (z + x)) 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1002
  using not_essential_shift_iff[of f 0 z] by (simp add: add_ac)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1003
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1004
lemma not_essential_holomorphic:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1005
  assumes "f holomorphic_on A" "x \<in> A" "open A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1006
  shows   "not_essential f x"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1007
  by (metis assms at_within_open continuous_on holomorphic_on_imp_continuous_on not_essential_def)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1008
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1009
lemma not_essential_analytic:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1010
  assumes "f analytic_on {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1011
  shows   "not_essential f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1012
  using analytic_at assms not_essential_holomorphic by blast
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1013
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1014
lemma not_essential_id [singularity_intros]: "not_essential (\<lambda>w. w) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1015
  by (simp add: not_essential_analytic)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1016
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1017
lemma is_pole_imp_not_essential [intro]: "is_pole f z \<Longrightarrow> not_essential f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1018
  by (auto simp: not_essential_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1019
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1020
lemma tendsto_imp_not_essential [intro]: "f \<midarrow>z\<rightarrow> c \<Longrightarrow> not_essential f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1021
  by (auto simp: not_essential_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1022
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1023
lemma eventually_not_pole:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1024
  assumes "isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1025
  shows   "eventually (\<lambda>w. \<not>is_pole f w) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1026
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1027
  from assms obtain r where "r > 0" and r: "f analytic_on ball z r - {z}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1028
    by (auto simp: isolated_singularity_at_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1029
  then have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1030
    by (intro eventually_at_in_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1031
  thus "eventually (\<lambda>w. \<not>is_pole f w) (at z)"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1032
    by (metis (no_types, lifting) analytic_at analytic_on_analytic_at eventually_mono not_is_pole_holomorphic r)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1033
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1034
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1035
lemma not_islimpt_poles:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1036
  assumes "isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1037
  shows   "\<not>z islimpt {w. is_pole f w}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1038
  using eventually_not_pole [OF assms]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1039
  by (auto simp: islimpt_conv_frequently_at frequently_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1040
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1041
lemma analytic_at_imp_no_pole: "f analytic_on {z} \<Longrightarrow> \<not>is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1042
  using analytic_at not_is_pole_holomorphic by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1043
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1044
lemma not_essential_const [singularity_intros]: "not_essential (\<lambda>_. c) z"
77277
c6b50597abbc More of Eberl's contributions: memomorphic functions
paulson <lp15@cam.ac.uk>
parents: 77228
diff changeset
  1045
  by blast
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1046
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1047
lemma not_essential_uminus [singularity_intros]:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1048
  assumes f_ness: "not_essential f z"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1049
  assumes f_iso: "isolated_singularity_at f z"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1050
  shows "not_essential (\<lambda>w. -f w) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1051
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1052
  have "not_essential (\<lambda>w. -1 * f w) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1053
    by (intro assms singularity_intros)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1054
  thus ?thesis by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1055
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1056
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1057
lemma isolated_singularity_at_analytic:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1058
  assumes "f analytic_on {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1059
  shows   "isolated_singularity_at f z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1060
  by (meson Diff_subset analytic_at assms holomorphic_on_subset isolated_singularity_at_holomorphic)
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1061
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1062
subsection \<open>The order of non-essential singularities (i.e. removable singularities or poles)\<close>
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1063
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1064
definition\<^marker>\<open>tag important\<close> zorder :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> int" where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1065
  "zorder f z = (THE n. (\<exists>h r. r>0 \<and> h holomorphic_on cball z r \<and> h z\<noteq>0
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1066
                   \<and> (\<forall>w\<in>cball z r - {z}. f w =  h w * (w-z) powi n
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1067
                   \<and> h w \<noteq>0)))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1068
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1069
definition\<^marker>\<open>tag important\<close> zor_poly
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1070
    :: "[complex \<Rightarrow> complex, complex] \<Rightarrow> complex \<Rightarrow> complex" where
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1071
  "zor_poly f z = (SOME h. \<exists>r. r > 0 \<and> h holomorphic_on cball z r \<and> h z \<noteq> 0
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1072
                   \<and> (\<forall>w\<in>cball z r - {z}. f w =  h w * (w-z) powi (zorder f z)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1073
                   \<and> h w \<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1074
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1075
lemma zorder_exist:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1076
  fixes f:: "complex \<Rightarrow> complex" and z::complex
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1077
  defines "n \<equiv> zorder f z" and "g \<equiv> zor_poly f z"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1078
  assumes f_iso: "isolated_singularity_at f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1079
      and f_ness: "not_essential f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1080
      and f_nconst: "\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1081
  shows "g z\<noteq>0 \<and> (\<exists>r. r>0 \<and> g holomorphic_on cball z r
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1082
    \<and> (\<forall>w\<in>cball z r - {z}. f w  = g w * (w-z) powi n  \<and> g w \<noteq>0))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1083
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1084
  define P where "P = (\<lambda>n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1085
          \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powi n \<and> g w\<noteq>0))"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1086
  have "\<exists>!k. \<exists>g r. P k g r"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1087
    using holomorphic_factor_puncture[OF assms(3-)] unfolding P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1088
  then have "\<exists>g r. P n g r"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1089
    unfolding n_def P_def zorder_def by (rule theI')
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1090
  then have "\<exists>r. P n g r"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1091
    unfolding P_def zor_poly_def g_def n_def by (rule someI_ex)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1092
  then obtain r1 where "P n g r1" 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1093
    by auto
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1094
  then show ?thesis 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1095
    unfolding P_def by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1096
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1097
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1098
lemma zorder_shift:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1099
  shows  "zorder f z = zorder (\<lambda>u. f (u + z)) 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1100
  unfolding zorder_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1101
  apply (rule arg_cong [of concl: The])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1102
  apply (auto simp: fun_eq_iff Ball_def dist_norm)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1103
  subgoal for x h r
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1104
    apply (rule_tac x="h o (+)z" in exI)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1105
    apply (rule_tac x="r" in exI)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1106
    apply (intro conjI holomorphic_on_compose holomorphic_intros)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1107
       apply (simp_all flip: cball_translation)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1108
    apply (simp add: add.commute)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1109
    done
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1110
  subgoal for x h r
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1111
    apply (rule_tac x="h o (\<lambda>u. u-z)" in exI)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1112
    apply (rule_tac x="r" in exI)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1113
    apply (intro conjI holomorphic_on_compose holomorphic_intros)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1114
       apply (simp_all flip: cball_translation_subtract)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1115
    by (metis diff_add_cancel eq_iff_diff_eq_0 norm_minus_commute)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1116
  done
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1117
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1118
lemma zorder_shift': "NO_MATCH 0 z \<Longrightarrow> zorder f z = zorder (\<lambda>u. f (u + z)) 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1119
  by (rule zorder_shift)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1120
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1121
lemma
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1122
  fixes f:: "complex \<Rightarrow> complex" and z::complex
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1123
  assumes f_iso: "isolated_singularity_at f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1124
      and f_ness: "not_essential f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1125
      and f_nconst: "\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1126
    shows zorder_inverse: "zorder (\<lambda>w. inverse (f w)) z = - zorder f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1127
      and zor_poly_inverse: "\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. inverse (f w)) z w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1128
                                                = inverse (zor_poly f z w)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1129
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1130
  define vf where "vf = (\<lambda>w. inverse (f w))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1131
  define fn vfn where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1132
    "fn = zorder f z"  and "vfn = zorder vf z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1133
  define fp vfp where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1134
    "fp = zor_poly f z" and "vfp = zor_poly vf z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1135
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1136
  obtain fr where [simp]: "fp z \<noteq> 0" and "fr > 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1137
          and fr: "fp holomorphic_on cball z fr"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1138
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w-z) powi fn \<and> fp w \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1139
    using zorder_exist[OF f_iso f_ness f_nconst,folded fn_def fp_def]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1140
    by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1141
  have fr_inverse: "vf w = (inverse (fp w)) * (w-z) powi (-fn)"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1142
        and fr_nz: "inverse (fp w) \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1143
    when "w\<in>ball z fr - {z}" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1144
  proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1145
    have "f w = fp w * (w-z) powi fn" "fp w \<noteq> 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1146
      using fr(2) that by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1147
    then show "vf w = (inverse (fp w)) * (w-z) powi (-fn)" "inverse (fp w)\<noteq>0"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1148
      by (simp_all add: power_int_minus vf_def)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1149
  qed
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1150
  obtain vfr where [simp]: "vfp z \<noteq> 0" and "vfr>0" and vfr: "vfp holomorphic_on cball z vfr"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1151
      "(\<forall>w\<in>cball z vfr - {z}. vf w = vfp w * (w-z) powi vfn \<and> vfp w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1152
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1153
    have "isolated_singularity_at vf z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1154
      using isolated_singularity_at_inverse[OF f_iso f_ness] unfolding vf_def .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1155
    moreover have "not_essential vf z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1156
      using not_essential_inverse[OF f_ness f_iso] unfolding vf_def .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1157
    moreover have "\<exists>\<^sub>F w in at z. vf w \<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1158
      using f_nconst unfolding vf_def by (auto elim: frequently_elim1)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1159
    ultimately show ?thesis using zorder_exist[of vf z, folded vfn_def vfp_def] that by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1160
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1161
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1162
  define r1 where "r1 = min fr vfr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1163
  have "r1>0" using \<open>fr>0\<close> \<open>vfr>0\<close> unfolding r1_def by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1164
  show "vfn = - fn"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1165
  proof (rule holomorphic_factor_unique)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1166
    have \<section>: "\<And>w. \<lbrakk>fp w = 0; dist z w < fr\<rbrakk> \<Longrightarrow> False"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1167
      using fr_nz by force
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1168
    then show "\<forall>w\<in>ball z r1 - {z}.
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1169
               vf w = vfp w * (w-z) powi vfn \<and>
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1170
               vfp w \<noteq> 0 \<and> vf w = inverse (fp w) * (w-z) powi (- fn) \<and>
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1171
               inverse (fp w) \<noteq> 0"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1172
      using fr_inverse r1_def vfr(2)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1173
      by (smt (verit) Diff_iff inverse_nonzero_iff_nonzero mem_ball mem_cball)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1174
    show "vfp holomorphic_on ball z r1"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1175
      using r1_def vfr(1) by auto
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1176
    show "(\<lambda>w. inverse (fp w)) holomorphic_on ball z r1"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1177
      by (metis \<section> ball_subset_cball fr(1) holomorphic_on_inverse holomorphic_on_subset mem_ball min.cobounded2 min.commute r1_def subset_ball)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1178
  qed (use \<open>r1>0\<close> in auto)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1179
  have "vfp w = inverse (fp w)" when "w\<in>ball z r1-{z}" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1180
  proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1181
    have "w \<in> ball z fr - {z}" "w \<in> cball z vfr - {z}"  "w\<noteq>z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1182
      using that unfolding r1_def by auto
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1183
    then show ?thesis
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1184
      by (metis \<open>vfn = - fn\<close> power_int_not_zero right_minus_eq  fr_inverse vfr(2)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1185
          vector_space_over_itself.scale_right_imp_eq) 
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1186
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1187
  then show "\<forall>\<^sub>Fw in (at z). vfp w = inverse (fp w)"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1188
    unfolding eventually_at by (metis DiffI dist_commute mem_ball singletonD \<open>r1>0\<close>)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1189
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1190
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1191
lemma zor_poly_shift:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1192
  assumes iso1: "isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1193
    and ness1: "not_essential f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1194
    and nzero1: "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1195
  shows "\<forall>\<^sub>F w in nhds z. zor_poly f z w = zor_poly (\<lambda>u. f (z + u)) 0 (w-z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1196
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1197
  obtain r1 where "r1>0" "zor_poly f z z \<noteq> 0" and
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1198
      holo1: "zor_poly f z holomorphic_on cball z r1" and
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1199
      rball1: "\<forall>w\<in>cball z r1 - {z}.
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1200
           f w = zor_poly f z w * (w-z) powi (zorder f z) \<and>
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1201
           zor_poly f z w \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1202
    using zorder_exist[OF iso1 ness1 nzero1] by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1203
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1204
  define ff where "ff=(\<lambda>u. f (z + u))"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1205
  have "isolated_singularity_at ff 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1206
    unfolding ff_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1207
    using iso1 isolated_singularity_at_shift_iff[of f 0 z]
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1208
    by (simp add: algebra_simps)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1209
  moreover have "not_essential ff 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1210
    unfolding ff_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1211
    using ness1 not_essential_shift_iff[of f 0 z]
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1212
    by (simp add: algebra_simps)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1213
  moreover have "\<exists>\<^sub>F w in at 0. ff w \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1214
    unfolding ff_def using nzero1
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1215
    by (smt (verit, ccfv_SIG) add.commute eventually_at_to_0
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1216
        eventually_mono not_frequently)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1217
  ultimately 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1218
  obtain r2 where "r2>0" "zor_poly ff 0 0 \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1219
          and holo2: "zor_poly ff 0 holomorphic_on cball 0 r2" 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1220
          and rball2: "\<forall>w\<in>cball 0 r2 - {0}.
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1221
               ff w = zor_poly ff 0 w * w powi (zorder ff 0) \<and> zor_poly ff 0 w \<noteq> 0"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1222
    using zorder_exist[of ff 0] by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1224
  define r where "r=min r1 r2"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1225
  have "r>0" using \<open>r1>0\<close> \<open>r2>0\<close> unfolding r_def by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1226
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1227
  have "zor_poly f z w = zor_poly ff 0 (w-z)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1228
    if "w\<in>ball z r - {z}" for w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1229
  proof -
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1230
    define n where "n \<equiv> zorder f z"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1231
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1232
    have "f w = zor_poly f z w * (w-z) powi n"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1233
      using n_def r_def rball1 that by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1234
    moreover have "f w = zor_poly ff 0 (w-z) * (w-z) powi n"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1235
    proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1236
      have "w-z\<in>cball 0 r2 - {0}"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1237
        using r_def that by (auto simp: dist_complex_def)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1238
      then have "ff (w-z) = zor_poly ff 0 (w-z) * (w-z) powi (zorder ff 0)"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1239
        using rball2 by blast
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1240
      moreover have "of_int (zorder ff 0) = n"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1241
        unfolding n_def ff_def by (simp add:zorder_shift' add.commute)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1242
      ultimately show ?thesis unfolding ff_def by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1243
    qed
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1244
    ultimately have "zor_poly f z w * (w-z) powi n = zor_poly ff 0 (w-z) * (w-z) powi n"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1245
      by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1246
    moreover have "(w-z) powi n \<noteq>0"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1247
      using that by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1248
    ultimately show ?thesis
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1249
      using mult_cancel_right by blast
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1250
  qed
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1251
  then have "\<forall>\<^sub>F w in at z. zor_poly f z w = zor_poly ff 0 (w-z)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1252
    unfolding eventually_at
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1253
    by (metis DiffI \<open>0 < r\<close> dist_commute mem_ball singletonD)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1254
  moreover have "isCont (zor_poly f z) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1255
    using holo1[THEN holomorphic_on_imp_continuous_on]
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1256
    by (simp add: \<open>0 < r1\<close> continuous_on_interior)
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1257
  moreover 
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1258
  have "isCont (zor_poly ff 0) 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1259
    using \<open>0 < r2\<close> continuous_on_interior holo2 holomorphic_on_imp_continuous_on
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1260
    by fastforce
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1261
  then have "isCont (\<lambda>w. zor_poly ff 0 (w-z)) z"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1262
      unfolding isCont_iff by simp
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1263
  ultimately show "\<forall>\<^sub>F w in nhds z. zor_poly f z w = zor_poly ff 0 (w-z)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1264
    by (elim at_within_isCont_imp_nhds;auto)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1265
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1266
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1267
lemma
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1268
  fixes f g:: "complex \<Rightarrow> complex" and z::complex
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1269
  assumes f_iso: "isolated_singularity_at f z" and g_iso: "isolated_singularity_at g z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1270
      and f_ness: "not_essential f z" and g_ness: "not_essential g z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1271
      and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1272
  shows zorder_times: "zorder (\<lambda>w. f w * g w) z = zorder f z + zorder g z" and
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1273
        zor_poly_times: "\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w * g w) z w
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1274
                                                  = zor_poly f z w *zor_poly g z w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1275
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1276
  define fg where "fg = (\<lambda>w. f w * g w)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1277
  define fn gn fgn where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1278
    "fn = zorder f z" and "gn = zorder g z" and "fgn = zorder fg z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1279
  define fp gp fgp where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1280
    "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1281
  have f_nconst: "\<exists>\<^sub>Fw in (at z). f w \<noteq> 0" and g_nconst: "\<exists>\<^sub>Fw in (at z).g w\<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1282
    using fg_nconst by (auto elim!:frequently_elim1)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1283
  obtain fr where [simp]: "fp z \<noteq> 0" and "fr > 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1284
          and fr: "fp holomorphic_on cball z fr"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1285
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w-z) powi fn \<and> fp w \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1286
    using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1287
  obtain gr where [simp]: "gp z \<noteq> 0" and "gr > 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1288
          and gr: "gp holomorphic_on cball z gr"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1289
                  "\<forall>w\<in>cball z gr - {z}. g w = gp w * (w-z) powi gn \<and> gp w \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1290
    using zorder_exist[OF g_iso g_ness g_nconst,folded gn_def gp_def] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1291
  define r1 where "r1=min fr gr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1292
  have "r1>0" unfolding r1_def using \<open>fr>0\<close> \<open>gr>0\<close> by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1293
  have fg_times: "fg w = (fp w * gp w) * (w-z) powi (fn+gn)" and fgp_nz: "fp w*gp w\<noteq>0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1294
    when "w\<in>ball z r1 - {z}" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1295
  proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1296
    have "f w = fp w * (w-z) powi fn" "fp w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1297
      using fr(2) r1_def that by auto
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1298
    moreover have "g w = gp w * (w-z) powi gn" "gp w \<noteq> 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1299
      using gr(2) that unfolding r1_def by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1300
    ultimately show "fg w = (fp w * gp w) * (w-z) powi (fn+gn)" "fp w*gp w\<noteq>0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1301
      using that unfolding fg_def by (auto simp add: power_int_add)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1302
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1303
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1304
  obtain fgr where [simp]: "fgp z \<noteq> 0" and "fgr > 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1305
          and fgr: "fgp holomorphic_on cball z fgr"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1306
                   "\<forall>w\<in>cball z fgr - {z}. fg w = fgp w * (w-z) powi fgn \<and> fgp w \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1307
  proof -
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1308
    have "isolated_singularity_at fg z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1309
      unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] .
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1310
    moreover have "not_essential fg z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1311
      by (simp add: f_iso f_ness fg_def g_iso g_ness not_essential_times)
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1312
    moreover have "\<exists>\<^sub>F w in at z. fg w \<noteq> 0"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1313
      using fg_def fg_nconst by blast
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1314
    ultimately show ?thesis 
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1315
      using that zorder_exist[of fg z] fgn_def fgp_def by fastforce
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1316
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1317
  define r2 where "r2 = min fgr r1"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1318
  have "r2>0" using \<open>r1>0\<close> \<open>fgr>0\<close> unfolding r2_def by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1319
  show "fgn = fn + gn "
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1320
  proof (rule holomorphic_factor_unique)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1321
    show "\<forall>w \<in> ball z r2 - {z}. fg w = fgp w * (w - z) powi fgn \<and> fgp w \<noteq> 0 \<and> fg w = fp w * gp w * (w - z) powi (fn + gn) \<and> fp w * gp w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1322
      using fg_times fgp_nz fgr(2) r2_def by fastforce
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1323
  next
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1324
    show "fgp holomorphic_on ball z r2"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1325
      using fgr(1) r2_def by auto
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1326
  next
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1327
    show "(\<lambda>w. fp w * gp w) holomorphic_on ball z r2"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1328
      by (metis ball_subset_cball fr(1) gr(1) holomorphic_on_mult holomorphic_on_subset
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1329
          min.cobounded1 min.cobounded2 r1_def r2_def subset_ball)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1330
  qed (auto simp add: \<open>0 < r2\<close>)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1331
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1332
  have "fgp w = fp w *gp w" when w: "w \<in> ball z r2-{z}" for w
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1333
  proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1334
    have "w \<in> ball z r1 - {z}" "w \<in> cball z fgr - {z}" "w\<noteq>z" 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1335
      using w unfolding r2_def by auto
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1336
    then show ?thesis
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1337
      by (metis \<open>fgn = fn + gn\<close> eq_iff_diff_eq_0 fg_times fgr(2) power_int_eq_0_iff
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1338
          mult_right_cancel)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1339
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1340
  then show "\<forall>\<^sub>Fw in (at z). fgp w = fp w * gp w"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1341
    using \<open>r2>0\<close> unfolding eventually_at by (auto simp add: dist_commute)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1342
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1343
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1344
lemma
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1345
  fixes f g:: "complex \<Rightarrow> complex" and z::complex
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1346
  assumes f_iso: "isolated_singularity_at f z" and g_iso: "isolated_singularity_at g z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1347
      and f_ness: "not_essential f z" and g_ness: "not_essential g z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1348
      and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1349
  shows zorder_divide: "zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z" and
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1350
        zor_poly_divide: "\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w / g w) z w
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1351
                                       = zor_poly f z w  / zor_poly g z w"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1352
proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1353
  have f_nconst: "\<exists>\<^sub>Fw in (at z). f w \<noteq> 0" and g_nconst: "\<exists>\<^sub>Fw in (at z).g w\<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1354
    using fg_nconst by (auto elim!:frequently_elim1)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1355
  define vg where "vg=(\<lambda>w. inverse (g w))"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1356
  have 1: "isolated_singularity_at vg z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1357
    by (simp add: g_iso g_ness isolated_singularity_at_inverse vg_def)
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1358
  moreover have 2: "not_essential vg z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1359
    by (simp add: g_iso g_ness not_essential_inverse vg_def)
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1360
  moreover have 3: "\<exists>\<^sub>F w in at z. f w * vg w \<noteq> 0"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1361
    using fg_nconst vg_def by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1362
  ultimately have "zorder (\<lambda>w. f w * vg w) z = zorder f z + zorder vg z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1363
    using zorder_times[OF f_iso _ f_ness] by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1364
  then show "zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1365
    using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1366
    by (auto simp add: field_simps)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1367
  have "\<forall>\<^sub>F w in at z. zor_poly (\<lambda>w. f w * vg w) z w = zor_poly f z w * zor_poly vg z w"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1368
    using zor_poly_times[OF f_iso _ f_ness,of vg] 1 2 3 by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1369
  then show "\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w / g w) z w = zor_poly f z w  / zor_poly g z w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1370
    using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1371
    by eventually_elim (auto simp add: field_simps)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1372
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1373
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1374
lemma zorder_exist_zero:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1375
  fixes f:: "complex \<Rightarrow> complex" and z::complex
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1376
  defines "n \<equiv> zorder f z" and "g \<equiv> zor_poly f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1377
  assumes  holo: "f holomorphic_on S" and "open S" "connected S" "z\<in>S"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1378
      and non_const: "\<exists>w\<in>S. f w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1379
  shows "(if f z=0 then n > 0 else n=0) \<and> (\<exists>r. r>0 \<and> cball z r \<subseteq> S \<and> g holomorphic_on cball z r
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1380
    \<and> (\<forall>w\<in>cball z r. f w  = g w * (w-z) ^ nat n  \<and> g w \<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1381
proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1382
  obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> S" "g holomorphic_on cball z r"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1383
            "(\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powi n \<and> g w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1384
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1385
    have "g z \<noteq> 0 \<and> (\<exists>r>0. g holomorphic_on cball z r
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1386
            \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powi n \<and> g w \<noteq> 0))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1387
    proof (rule zorder_exist[of f z,folded g_def n_def])
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1388
      show "isolated_singularity_at f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1389
        using \<open>open S\<close> \<open>z\<in>S\<close> holo holomorphic_on_imp_analytic_at isolated_singularity_at_analytic 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1390
        by force 
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1391
      show "not_essential f z" unfolding not_essential_def
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1392
        using \<open>open S\<close> \<open>z\<in>S\<close> at_within_open continuous_on holo holomorphic_on_imp_continuous_on
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1393
        by fastforce
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1394
      have "\<forall>\<^sub>F w in at z. f w \<noteq> 0 \<and> w\<in>S"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1395
        using assms(4,5,6) holo non_const non_zero_neighbour_alt by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1396
      then show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1397
        by (auto elim: eventually_frequentlyE)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1398
    qed
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1399
    then obtain r1 where "g z \<noteq> 0" "r1>0" and r1: "g holomorphic_on cball z r1"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1400
            "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w-z) powi n \<and> g w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1401
      by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1402
    obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> S"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1403
      using assms(4,6) open_contains_cball_eq by blast
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1404
    define r3 where "r3 \<equiv> min r1 r2"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1405
    have "r3>0" "cball z r3 \<subseteq> S" using \<open>r1>0\<close> r2 unfolding r3_def by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1406
    moreover have "g holomorphic_on cball z r3"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1407
      using r1(1) unfolding r3_def by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1408
    moreover have "(\<forall>w\<in>cball z r3 - {z}. f w = g w * (w-z) powi n \<and> g w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1409
      using r1(2) unfolding r3_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1410
    ultimately show ?thesis using that[of r3] \<open>g z\<noteq>0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1411
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1412
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1413
  have fz_lim: "f\<midarrow> z \<rightarrow> f z"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1414
    by (metis assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1415
  have gz_lim: "g \<midarrow>z\<rightarrow>g z"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1416
    using r
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1417
    by (meson Elementary_Metric_Spaces.open_ball analytic_at analytic_at_imp_isCont 
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1418
        ball_subset_cball centre_in_ball holomorphic_on_subset isContD)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1419
  have if_0: "if f z=0 then n > 0 else n=0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1420
  proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1421
    have "(\<lambda>w. g w * (w-z) powi n) \<midarrow>z\<rightarrow> f z"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1422
      using fz_lim Lim_transform_within_open[where s="ball z r"] r by fastforce
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1423
    then have "(\<lambda>w. (g w * (w-z) powi n) / g w) \<midarrow>z\<rightarrow> f z/g z"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1424
      using gz_lim \<open>g z \<noteq> 0\<close> tendsto_divide by blast
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1425
    then have powi_tendsto: "(\<lambda>w. (w-z) powi n) \<midarrow>z\<rightarrow> f z/g z"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1426
      using Lim_transform_within_open[where s="ball z r"] r by fastforce
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1427
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1428
    have ?thesis when "n\<ge>0" "f z=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1429
    proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1430
      have "(\<lambda>w. (w-z) ^ nat n) \<midarrow>z\<rightarrow> f z/g z"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1431
        using Lim_transform_within[OF powi_tendsto, where d=r]
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1432
        by (meson power_int_def r(1) that(1))
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1433
      then have *: "(\<lambda>w. (w-z) ^ nat n) \<midarrow>z\<rightarrow> 0" using \<open>f z=0\<close> by simp
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1434
      moreover have False when "n=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1435
      proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1436
        have "(\<lambda>w. (w-z) ^ nat n) \<midarrow>z\<rightarrow> 1"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1437
          using \<open>n=0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1438
        then show False using * using LIM_unique zero_neq_one by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1439
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1440
      ultimately show ?thesis using that by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1441
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1442
    moreover have ?thesis when "n\<ge>0" "f z\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1443
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1444
      have False when "n>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1445
      proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1446
        have "(\<lambda>w. (w-z) ^ nat n) \<midarrow>z\<rightarrow> f z/g z"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1447
          using Lim_transform_within[OF powi_tendsto, where d=r]
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1448
          by (meson \<open>0 \<le> n\<close> power_int_def r(1))
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1449
        moreover have "(\<lambda>w. (w-z) ^ nat n) \<midarrow>z\<rightarrow> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1450
          using \<open>n>0\<close> by (auto intro!:tendsto_eq_intros)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1451
        ultimately show False 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1452
          using \<open>f z\<noteq>0\<close> \<open>g z\<noteq>0\<close> LIM_unique divide_eq_0_iff by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1453
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1454
      then show ?thesis using that by force
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1455
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1456
    moreover have False when "n<0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1457
    proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1458
      have "(\<lambda>w. inverse ((w-z) ^ nat (- n))) \<midarrow>z\<rightarrow> f z/g z"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1459
        by (smt (verit) LIM_cong power_int_def power_inverse powi_tendsto that)
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1460
      moreover
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1461
      have "(\<lambda>w.((w-z) ^ nat (- n))) \<midarrow>z\<rightarrow> 0"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1462
        using that by (auto intro!:tendsto_eq_intros)
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1463
      ultimately
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1464
      have "(\<lambda>x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \<midarrow>z\<rightarrow> 0" 
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1465
        using tendsto_mult by fastforce
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1466
      then have "(\<lambda>x. 1::complex) \<midarrow>z\<rightarrow> 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1467
        using Lim_transform_within_open by fastforce
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1468
      then show False 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1469
        using LIM_const_eq by fastforce
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1470
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1471
    ultimately show ?thesis by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1472
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1473
  moreover have "f w  = g w * (w-z) ^ nat n  \<and> g w \<noteq>0" when "w\<in>cball z r" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1474
  proof (cases "w=z")
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1475
    case True
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1476
    then have "f \<midarrow>z\<rightarrow>f w"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1477
      using fz_lim by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1478
    then have "(\<lambda>w. g w * (w-z) ^ nat n) \<midarrow>z\<rightarrow>f w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1479
    proof (elim Lim_transform_within[OF _ \<open>r>0\<close>])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1480
      fix x assume "0 < dist x z" "dist x z < r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1481
      then have "x \<in> cball z r - {z}" "x\<noteq>z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1482
        unfolding cball_def by (auto simp add: dist_commute)
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1483
      then have "f x = g x * (x - z) powi n"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1484
        using r(4)[rule_format,of x] by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1485
      also have "... = g x * (x - z) ^ nat n"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1486
        by (smt (verit, best) if_0 int_nat_eq power_int_of_nat)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1487
      finally show "f x = g x * (x - z) ^ nat n" .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1488
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1489
    moreover have "(\<lambda>w. g w * (w-z) ^ nat n) \<midarrow>z\<rightarrow> g w * (w-z) ^ nat n"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1490
      using True by (auto intro!:tendsto_eq_intros gz_lim)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1491
    ultimately have "f w = g w * (w-z) ^ nat n" using LIM_unique by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1492
    then show ?thesis using \<open>g z\<noteq>0\<close> True by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1493
  next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1494
    case False
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1495
    then have "f w = g w * (w-z) powi n" "g w \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1496
      using r(4) that by auto
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1497
    then show ?thesis
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1498
      by (smt (verit, best) False if_0 int_nat_eq power_int_of_nat)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1499
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1500
  ultimately show ?thesis using r by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1501
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1502
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1503
lemma zorder_exist_pole:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1504
  fixes f:: "complex \<Rightarrow> complex" and z::complex
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1505
  defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1506
  assumes  holo: "f holomorphic_on S-{z}" and "open S" "z\<in>S" and "is_pole f z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1507
  shows "n < 0 \<and> g z\<noteq>0 \<and> (\<exists>r. r>0 \<and> cball z r \<subseteq> S \<and> g holomorphic_on cball z r
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1508
    \<and> (\<forall>w\<in>cball z r - {z}. f w  = g w / (w-z) ^ nat (- n) \<and> g w \<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1509
proof -
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1510
  obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> S" "g holomorphic_on cball z r"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1511
            "(\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powi n \<and> g w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1512
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1513
    have "g z \<noteq> 0 \<and> (\<exists>r>0. g holomorphic_on cball z r
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1514
            \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powi n \<and> g w \<noteq> 0))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1515
    proof (rule zorder_exist[of f z,folded g_def n_def])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1516
      show "isolated_singularity_at f z" unfolding isolated_singularity_at_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1517
        using holo assms(4,5)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1518
        by (metis analytic_on_holomorphic centre_in_ball insert_Diff openE open_delete subset_insert_iff)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1519
      show "not_essential f z" unfolding not_essential_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1520
        using assms(4,6) at_within_open continuous_on holo holomorphic_on_imp_continuous_on
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1521
        by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1522
      from non_zero_neighbour_pole[OF \<open>is_pole f z\<close>] show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1523
        by (auto elim: eventually_frequentlyE)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1524
    qed
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1525
    then obtain r1 where "g z \<noteq> 0" "r1>0" and r1: "g holomorphic_on cball z r1"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1526
            "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w-z) powi n \<and> g w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1527
      by auto
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1528
    obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> S"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1529
      using assms(4,5) open_contains_cball_eq by metis
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1530
    define r3 where "r3=min r1 r2"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1531
    have "r3>0" "cball z r3 \<subseteq> S" using \<open>r1>0\<close> r2 unfolding r3_def by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1532
    moreover have "g holomorphic_on cball z r3"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1533
      using r1(1) unfolding r3_def by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1534
    moreover have "(\<forall>w\<in>cball z r3 - {z}. f w = g w * (w-z) powi n \<and> g w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1535
      using r1(2) unfolding r3_def by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1536
    ultimately show ?thesis 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1537
      using that[of r3] \<open>g z\<noteq>0\<close> by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1538
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1539
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1540
  have "n<0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1541
  proof (rule ccontr)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1542
    assume " \<not> n < 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1543
    define c where "c=(if n=0 then g z else 0)"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1544
    have [simp]: "g \<midarrow>z\<rightarrow> g z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1545
      using r
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1546
      by (metis centre_in_ball continuous_on_interior holomorphic_on_imp_continuous_on
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1547
          interior_cball isContD)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1548
    have "\<forall>x \<in> ball z r. x \<noteq> z \<longrightarrow> f x = g x * (x - z) ^ nat n"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1549
      by (simp add: \<open>\<not> n < 0\<close> linorder_not_le power_int_def r)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1550
    then have "\<forall>\<^sub>F x in at z. f x = g x * (x - z) ^ nat n"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1551
      using centre_in_ball eventually_at_topological r(1) by blast
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1552
    moreover have "(\<lambda>x. g x * (x - z) ^ nat n) \<midarrow>z\<rightarrow> c"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1553
    proof (cases "n=0")
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1554
      case True
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1555
      then show ?thesis unfolding c_def by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1556
    next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1557
      case False
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1558
      then have "(\<lambda>x. (x - z) ^ nat n) \<midarrow>z\<rightarrow> 0" using \<open>\<not> n < 0\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1559
        by (auto intro!:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1560
      from tendsto_mult[OF _ this,of g "g z",simplified]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1561
      show ?thesis unfolding c_def using False by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1562
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1563
    ultimately have "f \<midarrow>z\<rightarrow>c" using tendsto_cong by fast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1564
    then show False using \<open>is_pole f z\<close> at_neq_bot not_tendsto_and_filterlim_at_infinity
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1565
      unfolding is_pole_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1566
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1567
  moreover have "\<forall>w\<in>cball z r - {z}. f w  = g w / (w-z) ^ nat (- n) \<and> g w \<noteq>0"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1568
    using r(4) \<open>n<0\<close>
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1569
    by (smt (verit) inverse_eq_divide mult.right_neutral power_int_def power_inverse times_divide_eq_right)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1570
  ultimately show ?thesis 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1571
    using r \<open>g z\<noteq>0\<close> by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1572
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1573
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1574
lemma zorder_eqI:
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1575
  assumes "open S" "z \<in> S" "g holomorphic_on S" "g z \<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1576
  assumes fg_eq: "\<And>w. \<lbrakk>w \<in> S;w\<noteq>z\<rbrakk> \<Longrightarrow> f w = g w * (w-z) powi n"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1577
  shows   "zorder f z = n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1578
proof -
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1579
  have "continuous_on S g" by (rule holomorphic_on_imp_continuous_on) fact
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1580
  moreover have "open (-{0::complex})" by auto
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1581
  ultimately have "open ((g -` (-{0})) \<inter> S)"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1582
    unfolding continuous_on_open_vimage[OF \<open>open S\<close>] by blast
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1583
  moreover from assms have "z \<in> (g -` (-{0})) \<inter> S" by auto
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1584
  ultimately obtain r where r: "r > 0" "cball z r \<subseteq>  S \<inter> (g -` (-{0}))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1585
    unfolding open_contains_cball by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1586
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1587
  let ?gg= "(\<lambda>w. g w * (w-z) powi n)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1588
  define P where "P = (\<lambda>n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1589
          \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powi n \<and> g w\<noteq>0))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1590
  have "P n g r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1591
    unfolding P_def using r assms(3,4,5) by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1592
  then have "\<exists>g r. P n g r" by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1593
  moreover have unique: "\<exists>!n. \<exists>g r. P n g r" unfolding P_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1594
  proof (rule holomorphic_factor_puncture)
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1595
    have "ball z r-{z} \<subseteq> S" using r using ball_subset_cball by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1596
    then have "?gg holomorphic_on ball z r-{z}"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1597
      using \<open>g holomorphic_on S\<close> r by (auto intro!: holomorphic_intros)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1598
    then have "f holomorphic_on ball z r - {z}"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1599
      by (smt (verit, best) DiffD2 \<open>ball z r-{z} \<subseteq> S\<close> fg_eq holomorphic_cong singleton_iff subset_iff)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1600
    then show "isolated_singularity_at f z" unfolding isolated_singularity_at_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1601
      using analytic_on_open open_delete r(1) by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1602
  next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1603
    have "not_essential ?gg z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1604
    proof (intro singularity_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1605
      show "not_essential g z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1606
        by (meson \<open>continuous_on S g\<close> assms continuous_on_eq_continuous_at
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1607
            isCont_def not_essential_def)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1608
      show " \<forall>\<^sub>F w in at z. w - z \<noteq> 0" by (simp add: eventually_at_filter)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1609
      then show "LIM w at z. w - z :> at 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1610
        unfolding filterlim_at by (auto intro: tendsto_eq_intros)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1611
      show "isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1612
        by (meson Diff_subset open_ball analytic_on_holomorphic
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1613
            assms holomorphic_on_subset isolated_singularity_at_def openE)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1614
    qed
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1615
    moreover
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1616
    have "\<forall>\<^sub>F w in at z. g w * (w-z) powi n = f w"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1617
      unfolding eventually_at_topological using assms fg_eq by force
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1618
    ultimately show "not_essential f z"
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1619
      using not_essential_transform by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1620
    show "\<exists>\<^sub>F w in at z. f w \<noteq> 0" unfolding frequently_at
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1621
    proof (intro strip)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1622
      fix d::real assume "0 < d"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1623
      define z' where "z' \<equiv> z+min d r / 2"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1624
      have "z' \<noteq> z" " dist z' z < d "
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1625
        unfolding z'_def using \<open>d>0\<close> \<open>r>0\<close> by (auto simp add: dist_norm)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1626
      moreover have "f z' \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1627
      proof (subst fg_eq[OF _ \<open>z'\<noteq>z\<close>])
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1628
        have "z' \<in> cball z r"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1629
          unfolding z'_def using \<open>r>0\<close> \<open>d>0\<close> by (auto simp add: dist_norm)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1630
        then show "z' \<in> S" using r(2) by blast
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1631
        show "g z' * (z' - z) powi n \<noteq> 0"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1632
          using P_def \<open>P n g r\<close> \<open>z' \<in> cball z r\<close> \<open>z' \<noteq> z\<close> by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1633
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1634
      ultimately show "\<exists>x\<in>UNIV. x \<noteq> z \<and> dist x z < d \<and> f x \<noteq> 0" by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1635
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1636
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1637
  ultimately have "(THE n. \<exists>g r. P n g r) = n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1638
    by (rule_tac the1_equality)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1639
  then show ?thesis unfolding zorder_def P_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1640
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1641
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1642
lemma simple_zeroI:
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1643
  assumes "open S" "z \<in> S" "g holomorphic_on S" "g z \<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1644
  assumes "\<And>w. w \<in> S \<Longrightarrow> f w = g w * (w-z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1645
  shows   "zorder f z = 1"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1646
  using assms zorder_eqI by force
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1647
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1648
lemma higher_deriv_power:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1649
  shows   "(deriv ^^ j) (\<lambda>w. (w-z) ^ n) w =
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1650
             pochhammer (of_nat (Suc n - j)) j * (w-z) ^ (n - j)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1651
proof (induction j arbitrary: w)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1652
  case 0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1653
  thus ?case by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1654
next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1655
  case (Suc j w)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1656
  have "(deriv ^^ Suc j) (\<lambda>w. (w-z) ^ n) w = deriv ((deriv ^^ j) (\<lambda>w. (w-z) ^ n)) w"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1657
    by simp
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1658
  also have "(deriv ^^ j) (\<lambda>w. (w-z) ^ n) =
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1659
               (\<lambda>w. pochhammer (of_nat (Suc n - j)) j * (w-z) ^ (n - j))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1660
    using Suc by (intro Suc.IH ext)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1661
  also {
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1662
    have "(\<dots> has_field_derivative of_nat (n - j) *
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1663
               pochhammer (of_nat (Suc n - j)) j * (w-z) ^ (n - Suc j)) (at w)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1664
      using Suc.prems by (auto intro!: derivative_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1665
    also have "of_nat (n - j) * pochhammer (of_nat (Suc n - j)) j =
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1666
                 pochhammer (of_nat (Suc n - Suc j)) (Suc j)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1667
      by (cases "Suc j \<le> n", subst pochhammer_rec)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1668
         (use Suc.prems in \<open>simp_all add: algebra_simps Suc_diff_le pochhammer_0_left\<close>)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1669
    finally have "deriv (\<lambda>w. pochhammer (of_nat (Suc n - j)) j * (w-z) ^ (n - j)) w =
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1670
                    \<dots> * (w-z) ^ (n - Suc j)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1671
      by (rule DERIV_imp_deriv)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1672
  }
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1673
  finally show ?case .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1674
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1675
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1676
lemma zorder_zero_eqI:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1677
  assumes  f_holo: "f holomorphic_on S" and "open S" "z \<in> S"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1678
  assumes zero: "\<And>i. i < nat n \<Longrightarrow> (deriv ^^ i) f z = 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1679
  assumes nz: "(deriv ^^ nat n) f z \<noteq> 0" and "n\<ge>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1680
  shows   "zorder f z = n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1681
proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1682
  obtain r where [simp]: "r>0" and "ball z r \<subseteq> S"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1683
    using \<open>open S\<close> \<open>z\<in>S\<close> openE by blast
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1684
  have nz': "\<exists>w\<in>ball z r. f w \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1685
  proof (rule ccontr)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1686
    assume "\<not> (\<exists>w\<in>ball z r. f w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1687
    then have "eventually (\<lambda>u. f u = 0) (nhds z)"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1688
      using open_ball \<open>0 < r\<close> centre_in_ball eventually_nhds by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1689
    then have "(deriv ^^ nat n) f z = (deriv ^^ nat n) (\<lambda>_. 0) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1690
      by (intro higher_deriv_cong_ev) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1691
    also have "(deriv ^^ nat n) (\<lambda>_. 0) z = 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1692
      by (induction n) simp_all
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1693
    finally show False using nz by contradiction
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1694
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1695
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1696
  define zn g where "zn = zorder f z" and "g = zor_poly f z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1697
  obtain e where e_if: "if f z = 0 then 0 < zn else zn = 0" and
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1698
            [simp]: "e>0" and "cball z e \<subseteq> ball z r" and
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1699
            g_holo: "g holomorphic_on cball z e" and
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1700
            e_fac: "(\<forall>w\<in>cball z e. f w = g w * (w-z) ^ nat zn \<and> g w \<noteq> 0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1701
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1702
    have "f holomorphic_on ball z r"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1703
      using f_holo \<open>ball z r \<subseteq> S\<close> by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1704
    from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def]
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1705
    show thesis by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1706
  qed
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1707
  then obtain "zn \<ge> 0" "g z \<noteq> 0"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1708
    by (metis centre_in_cball less_le_not_le order_refl)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1709
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1710
  define A where "A \<equiv> (\<lambda>i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1711
  have deriv_A: "(deriv ^^ i) f z = (if zn \<le> int i then A i else 0)" for i
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1712
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1713
    have "eventually (\<lambda>w. w \<in> ball z e) (nhds z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1714
      using \<open>cball z e \<subseteq> ball z r\<close> \<open>e>0\<close> by (intro eventually_nhds_in_open) auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1715
    hence "eventually (\<lambda>w. f w = (w-z) ^ (nat zn) * g w) (nhds z)"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1716
      using e_fac eventually_mono by fastforce
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1717
    hence "(deriv ^^ i) f z = (deriv ^^ i) (\<lambda>w. (w-z) ^ nat zn * g w) z"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1718
      by (intro higher_deriv_cong_ev) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1719
    also have "\<dots> = (\<Sum>j=0..i. of_nat (i choose j) *
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1720
                       (deriv ^^ j) (\<lambda>w. (w-z) ^ nat zn) z * (deriv ^^ (i - j)) g z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1721
      using g_holo \<open>e>0\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1722
      by (intro higher_deriv_mult[of _ "ball z e"]) (auto intro!: holomorphic_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1723
    also have "\<dots> = (\<Sum>j=0..i. if j = nat zn then
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1724
                    of_nat (i choose nat zn) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z else 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1725
    proof (intro sum.cong refl, goal_cases)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1726
      case (1 j)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1727
      have "(deriv ^^ j) (\<lambda>w. (w-z) ^ nat zn) z =
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1728
              pochhammer (of_nat (Suc (nat zn) - j)) j * 0 ^ (nat zn - j)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1729
        by (subst higher_deriv_power) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1730
      also have "\<dots> = (if j = nat zn then fact j else 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1731
        by (auto simp: not_less pochhammer_0_left pochhammer_fact)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1732
      also have "of_nat (i choose j) * \<dots> * (deriv ^^ (i - j)) g z =
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1733
                   (if j = nat zn then of_nat (i choose (nat zn)) * fact (nat zn)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1734
                        * (deriv ^^ (i - nat zn)) g z else 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1735
        by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1736
      finally show ?case .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1737
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1738
    also have "\<dots> = (if i \<ge> zn then A i else 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1739
      by (auto simp: A_def)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1740
    finally show "(deriv ^^ i) f z = \<dots>" .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1741
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1742
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1743
  have False when "n<zn"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1744
    using deriv_A[of "nat n"] that \<open>n\<ge>0\<close> by (simp add: nz) 
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1745
  moreover have "n\<le>zn"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1746
  proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1747
    have "g z \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1748
      by (simp add: \<open>g z \<noteq> 0\<close>)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1749
    then have "(deriv ^^ nat zn) f z \<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1750
      using deriv_A[of "nat zn"] by(auto simp add: A_def)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1751
    then have "nat zn \<ge> nat n" using zero[of "nat zn"] by linarith
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1752
    moreover have "zn\<ge>0" using e_if by (auto split: if_splits)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1753
    ultimately show ?thesis using nat_le_eq_zle by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1754
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1755
  ultimately show ?thesis unfolding zn_def by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1756
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1757
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1758
lemma
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1759
  assumes "eventually (\<lambda>z. f z = g z) (at z)" "z = z'"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1760
  shows zorder_cong: "zorder f z = zorder g z'" and zor_poly_cong: "zor_poly f z = zor_poly g z'"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1761
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1762
  define P where "P = (\<lambda>ff n h r. 0 < r \<and> h holomorphic_on cball z r \<and> h z\<noteq>0
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1763
                    \<and> (\<forall>w\<in>cball z r - {z}. ff w = h w * (w-z) powi n \<and> h w\<noteq>0))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1764
  have "(\<exists>r. P f n h r) = (\<exists>r. P g n h r)" for n h
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1765
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1766
    have *: "\<exists>r. P g n h r" if "\<exists>r. P f n h r" and "eventually (\<lambda>x. f x = g x) (at z)" for f g
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1767
    proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1768
      from that(1) obtain r1 where r1_P: "P f n h r1" by auto
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1769
      from that(2) obtain r2 where "r2>0" and r2_dist: "\<forall>x. x \<noteq> z \<and> dist x z \<le> r2 \<longrightarrow> f x = g x"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1770
        unfolding eventually_at_le by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1771
      define r where "r=min r1 r2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1772
      have "r>0" "h z\<noteq>0" using r1_P \<open>r2>0\<close> unfolding r_def P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1773
      moreover have "h holomorphic_on cball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1774
        using r1_P unfolding P_def r_def by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1775
      moreover have "g w = h w * (w-z) powi n \<and> h w \<noteq> 0" when "w\<in>cball z r - {z}" for w
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1776
      proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1777
        have "f w = h w * (w-z) powi n \<and> h w \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1778
          using r1_P that unfolding P_def r_def by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1779
        moreover have "f w=g w"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1780
          using r2_dist that by (simp add: dist_commute r_def)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1781
        ultimately show ?thesis by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1782
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1783
      ultimately show ?thesis unfolding P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1784
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1785
    from assms have eq': "eventually (\<lambda>z. g z = f z) (at z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1786
      by (simp add: eq_commute)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1787
    show ?thesis
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1788
      using "*" assms(1) eq' by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1789
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1790
  then show "zorder f z = zorder g z'" "zor_poly f z = zor_poly g z'"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1791
      using \<open>z=z'\<close> unfolding P_def zorder_def zor_poly_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1792
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1793
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1794
lemma zorder_times_analytic':
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1795
  assumes "isolated_singularity_at f z" "not_essential f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1796
  assumes "g analytic_on {z}" "frequently (\<lambda>z. f z * g z \<noteq> 0) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1797
  shows   "zorder (\<lambda>x. f x * g x) z = zorder f z + zorder g z"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1798
  using assms isolated_singularity_at_analytic not_essential_analytic zorder_times by blast
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1799
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1800
lemma zorder_cmult:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1801
  assumes "c \<noteq> 0"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1802
  shows   "zorder (\<lambda>z. c * f z) z = zorder f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1803
proof -
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1804
  define P where
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1805
    "P = (\<lambda>f n h r. 0 < r \<and> h holomorphic_on cball z r \<and>
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1806
                    h z \<noteq> 0 \<and> (\<forall>w\<in>cball z r - {z}. f w = h w * (w-z) powi n \<and> h w \<noteq> 0))"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1807
  have *: "P (\<lambda>x. c * f x) n (\<lambda>x. c * h x) r" 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1808
    if "P f n h r" "c \<noteq> 0" for f n h r c
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1809
    using that unfolding P_def by (auto intro!: holomorphic_intros)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1810
  have "(\<exists>h r. P (\<lambda>x. c * f x) n h r) \<longleftrightarrow> (\<exists>h r. P f n h r)" for n
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1811
    using *[of f n _ _ c] *[of "\<lambda>x. c * f x" n _ _ "inverse c"] \<open>c \<noteq> 0\<close>
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1812
    by (fastforce simp: field_simps)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1813
  hence "(THE n. \<exists>h r. P (\<lambda>x. c * f x) n h r) = (THE n. \<exists>h r. P f n h r)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1814
    by simp
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1815
  thus ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1816
    by (simp add: zorder_def P_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1817
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  1818
79945
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78517
diff changeset
  1819
lemma zorder_uminus [simp]: "zorder (\<lambda>z. -f z) z = zorder f z"
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78517
diff changeset
  1820
  using zorder_cmult[of "-1" f] by simp
ca004ccf2352 New material from a variety of sources (including AFP)
paulson <lp15@cam.ac.uk>
parents: 78517
diff changeset
  1821
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1822
lemma zorder_nonzero_div_power:
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1823
  assumes sz: "open S" "z \<in> S" "f holomorphic_on S" "f z \<noteq> 0" and "n > 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1824
  shows  "zorder (\<lambda>w. f w / (w-z) ^ n) z = - n"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1825
  by (intro zorder_eqI [OF sz]) (simp add: inverse_eq_divide power_int_minus)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1826
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1827
lemma zor_poly_eq:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1828
  assumes "isolated_singularity_at f z" "not_essential f z" "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1829
  shows "eventually (\<lambda>w. zor_poly f z w = f w * (w-z) powi - zorder f z) (at z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1830
proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1831
  obtain r where r: "r>0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1832
       "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w * (w-z) powi (zorder f z))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1833
    using zorder_exist[OF assms] by blast
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1834
  then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w * (w-z) powi - zorder f z"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1835
    by (auto simp: field_simps power_int_minus)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1836
  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1837
    using r eventually_at_ball'[of r z UNIV] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1838
  thus ?thesis by eventually_elim (insert *, auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1839
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1840
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1841
lemma zor_poly_zero_eq:
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  1842
  assumes "f holomorphic_on S" "open S" "connected S" "z \<in> S" "\<exists>w\<in>S. f w \<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1843
  shows "eventually (\<lambda>w. zor_poly f z w = f w / (w-z) ^ nat (zorder f z)) (at z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1844
proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1845
  obtain r where r: "r>0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1846
       "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w * (w-z) ^ nat (zorder f z))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1847
    using zorder_exist_zero[OF assms] by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1848
  then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w / (w-z) ^ nat (zorder f z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1849
    by (auto simp: field_simps powr_minus)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1850
  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1851
    using r eventually_at_ball'[of r z UNIV] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1852
  thus ?thesis by eventually_elim (insert *, auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1853
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1854
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1855
lemma zor_poly_pole_eq:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1856
  assumes f_iso: "isolated_singularity_at f z" "is_pole f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1857
  shows "eventually (\<lambda>w. zor_poly f z w = f w * (w-z) ^ nat (- zorder f z)) (at z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1858
proof -
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1859
  obtain e where [simp]: "e>0" and f_holo: "f holomorphic_on ball z e - {z}"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1860
    using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1861
  obtain r where r: "r>0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1862
       "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w / (w-z) ^ nat (- zorder f z))"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1863
    using zorder_exist_pole[OF f_holo,simplified,OF \<open>is_pole f z\<close>] by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1864
  then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w * (w-z) ^ nat (- zorder f z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1865
    by (auto simp: field_simps)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1866
  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1867
    using r eventually_at_ball'[of r z UNIV] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1868
  thus ?thesis by eventually_elim (insert *, auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1869
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1870
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1871
lemma zor_poly_eqI:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1872
  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1873
  defines "n \<equiv> zorder f z0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1874
  assumes "isolated_singularity_at f z0" "not_essential f z0" "\<exists>\<^sub>F w in at z0. f w \<noteq> 0"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1875
  assumes lim: "((\<lambda>x. f (g x) * (g x - z0) powi - n) \<longlongrightarrow> c) F"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1876
  assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1877
  shows   "zor_poly f z0 z0 = c"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1878
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1879
  from zorder_exist[OF assms(2-4)] obtain r where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1880
    r: "r > 0" "zor_poly f z0 holomorphic_on cball z0 r"
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1881
       "\<And>w. w \<in> cball z0 r - {z0} \<Longrightarrow> f w = zor_poly f z0 w * (w - z0) powi n"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1882
    unfolding n_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1883
  from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1884
    using eventually_at_ball'[of r z0 UNIV] by auto
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1885
  hence "eventually (\<lambda>w. zor_poly f z0 w = f w * (w - z0) powi - n) (at z0)"
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1886
    by eventually_elim (insert r, auto simp: field_simps power_int_minus)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1887
  moreover have "continuous_on (ball z0 r) (zor_poly f z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1888
    using r by (intro holomorphic_on_imp_continuous_on) auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1889
  with r have "isCont (zor_poly f z0) z0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1890
    by (auto simp: continuous_on_eq_continuous_at)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1891
  hence "(zor_poly f z0 \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1892
    unfolding isCont_def .
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1893
  ultimately have "((\<lambda>w. f w * (w - z0) powi - n) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1894
    by (blast intro: Lim_transform_eventually)
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  1895
  hence "((\<lambda>x. f (g x) * (g x - z0) powi - n) \<longlongrightarrow> zor_poly f z0 z0) F"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1896
    by (rule filterlim_compose[OF _ g])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1897
  from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1898
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1899
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1900
lemma zor_poly_zero_eqI:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1901
  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1902
  defines "n \<equiv> zorder f z0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1903
  assumes "f holomorphic_on A" "open A" "connected A" "z0 \<in> A" "\<exists>z\<in>A. f z \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1904
  assumes lim: "((\<lambda>x. f (g x) / (g x - z0) ^ nat n) \<longlongrightarrow> c) F"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1905
  assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1906
  shows   "zor_poly f z0 z0 = c"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1907
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1908
  from zorder_exist_zero[OF assms(2-6)] obtain r where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1909
    r: "r > 0" "cball z0 r \<subseteq> A" "zor_poly f z0 holomorphic_on cball z0 r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1910
       "\<And>w. w \<in> cball z0 r \<Longrightarrow> f w = zor_poly f z0 w * (w - z0) ^ nat n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1911
    unfolding n_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1912
  from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1913
    using eventually_at_ball'[of r z0 UNIV] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1914
  hence "eventually (\<lambda>w. zor_poly f z0 w = f w / (w - z0) ^ nat n) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1915
    by eventually_elim (insert r, auto simp: field_simps)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1916
  moreover have "continuous_on (ball z0 r) (zor_poly f z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1917
    using r by (intro holomorphic_on_imp_continuous_on) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1918
  with r(1,2) have "isCont (zor_poly f z0) z0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1919
    by (auto simp: continuous_on_eq_continuous_at)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1920
  hence "(zor_poly f z0 \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1921
    unfolding isCont_def .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1922
  ultimately have "((\<lambda>w. f w / (w - z0) ^ nat n) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1923
    by (blast intro: Lim_transform_eventually)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1924
  hence "((\<lambda>x. f (g x) / (g x - z0) ^ nat n) \<longlongrightarrow> zor_poly f z0 z0) F"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1925
    by (rule filterlim_compose[OF _ g])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1926
  from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1927
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1928
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1929
lemma zor_poly_pole_eqI:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1930
  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1931
  defines "n \<equiv> zorder f z0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1932
  assumes f_iso: "isolated_singularity_at f z0" and "is_pole f z0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1933
  assumes lim: "((\<lambda>x. f (g x) * (g x - z0) ^ nat (-n)) \<longlongrightarrow> c) F"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1934
  assumes g: "filterlim g (at z0) F" and "F \<noteq> bot"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1935
  shows   "zor_poly f z0 z0 = c"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1936
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1937
  obtain r where r: "r > 0"  "zor_poly f z0 holomorphic_on cball z0 r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1938
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1939
    have "\<exists>\<^sub>F w in at z0. f w \<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1940
      using non_zero_neighbour_pole[OF \<open>is_pole f z0\<close>] 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1941
      by (auto elim: eventually_frequentlyE)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1942
    moreover have "not_essential f z0" 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1943
      unfolding not_essential_def using \<open>is_pole f z0\<close> by simp
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1944
    ultimately show ?thesis 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1945
      using that zorder_exist[OF f_iso,folded n_def] by auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1946
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1947
  from r(1) have "eventually (\<lambda>w. w \<in> ball z0 r \<and> w \<noteq> z0) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1948
    using eventually_at_ball'[of r z0 UNIV] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1949
  have "eventually (\<lambda>w. zor_poly f z0 w = f w * (w - z0) ^ nat (-n)) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1950
    using zor_poly_pole_eq[OF f_iso \<open>is_pole f z0\<close>] unfolding n_def .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1951
  moreover have "continuous_on (ball z0 r) (zor_poly f z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1952
    using r by (intro holomorphic_on_imp_continuous_on) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1953
  with r(1,2) have "isCont (zor_poly f z0) z0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1954
    by (auto simp: continuous_on_eq_continuous_at)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1955
  hence "(zor_poly f z0 \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1956
    unfolding isCont_def .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1957
  ultimately have "((\<lambda>w. f w * (w - z0) ^ nat (-n)) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1958
    by (blast intro: Lim_transform_eventually)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1959
  hence "((\<lambda>x. f (g x) * (g x - z0) ^ nat (-n)) \<longlongrightarrow> zor_poly f z0 z0) F"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1960
    by (rule filterlim_compose[OF _ g])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1961
  from tendsto_unique[OF \<open>F \<noteq> bot\<close> this lim] show ?thesis .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1962
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1963
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1964
lemma
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1965
  assumes "is_pole f (x :: complex)" "open A" "x \<in> A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1966
  assumes "\<And>y. y \<in> A - {x} \<Longrightarrow> (f has_field_derivative f' y) (at y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1967
  shows   is_pole_deriv': "is_pole f' x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1968
    and   zorder_deriv':  "zorder f' x = zorder f x - 1"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1969
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1970
  have holo: "f holomorphic_on A - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1971
    using assms by (subst holomorphic_on_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1972
  obtain r where r: "r > 0" "ball x r \<subseteq> A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1973
    using assms(2,3) openE by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1974
  moreover have "open (ball x r - {x})"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1975
    by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1976
  ultimately have "isolated_singularity_at f x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1977
    by (auto simp: isolated_singularity_at_def analytic_on_open
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1978
             intro!: exI[of _ r] holomorphic_on_subset[OF holo])
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1979
  hence ev: "\<forall>\<^sub>F w in at x. zor_poly f x w = f w * (w-x) ^ nat (- zorder f x)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1980
    using \<open>is_pole f x\<close> zor_poly_pole_eq by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1981
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1982
  define P where "P = zor_poly f x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1983
  define n where "n = nat (-zorder f x)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1984
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1985
  obtain r where r: "r > 0" "cball x r \<subseteq> A" "P holomorphic_on cball x r" "zorder f x < 0" "P x \<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1986
    "\<forall>w\<in>cball x r - {x}. f w = P w / (w-x) ^ n \<and> P w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1987
    using P_def assms holo n_def zorder_exist_pole by blast
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1988
  have n: "n > 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1989
    using r(4) by (auto simp: n_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1990
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1991
  have [derivative_intros]: "(P has_field_derivative deriv P w) (at w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1992
    if "w \<in> ball x r" for w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1993
    using that by (intro holomorphic_derivI[OF holomorphic_on_subset[OF r(3), of "ball x r"]]) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1994
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  1995
  define D where "D = (\<lambda>w. (deriv P w * (w-x) - of_nat n * P w) / (w-x) ^ (n + 1))"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1996
  define n' where "n' = n - 1"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1997
  have n': "n = Suc n'"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1998
    using n by (simp add: n'_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1999
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2000
  have "eventually (\<lambda>w. w \<in> ball x r) (nhds x)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2001
    using \<open>r > 0\<close> by (intro eventually_nhds_in_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2002
  hence ev'': "eventually (\<lambda>w. w \<in> ball x r - {x}) (at x)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2003
    by (auto simp: eventually_at_filter elim: eventually_mono)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2004
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2005
  {
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2006
    fix w assume w: "w \<in> ball x r - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2007
    have ev': "eventually (\<lambda>w. w \<in> ball x r - {x}) (nhds w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2008
      using w by (intro eventually_nhds_in_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2009
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2010
    have \<section>: "(deriv P w * (w-x) ^ n - P w * (n * (w-x) ^ (n-1))) / ((w-x) ^ n * (w-x) ^ n) = D w"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2011
      using w n' by (simp add: divide_simps D_def) (simp add: algebra_simps)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2012
    have "((\<lambda>w. P w / (w-x) ^ n) has_field_derivative D w) (at w)"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2013
      by (rule derivative_eq_intros refl | use w \<section> in force)+
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2014
    also have "?this \<longleftrightarrow> (f has_field_derivative D w) (at w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2015
      using r by (intro has_field_derivative_cong_ev refl eventually_mono[OF ev']) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2016
    finally have "(f has_field_derivative D w) (at w)" .
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2017
    moreover have "(f has_field_derivative f' w) (at w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2018
      using w r by (intro assms) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2019
    ultimately have "D w = f' w"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2020
      using DERIV_unique by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2021
  } note D_eq = this
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2022
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2023
  have "is_pole D x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2024
    unfolding D_def using n \<open>r > 0\<close> \<open>P x \<noteq> 0\<close>
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2025
    by (intro is_pole_basic[where A = "ball x r"] holomorphic_intros holomorphic_on_subset[OF r(3)]) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2026
  also have "?this \<longleftrightarrow> is_pole f' x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2027
    by (intro is_pole_cong eventually_mono[OF ev''] D_eq) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2028
  finally show "is_pole f' x" .
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2029
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2030
  have "zorder f' x = -int (Suc n)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2031
  proof (rule zorder_eqI)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2032
    show "open (ball x r)" "x \<in> ball x r"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2033
      using \<open>r > 0\<close> by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2034
    show "f' w = (deriv P w * (w-x) - of_nat n * P w) * (w-x) powi (- int (Suc n))"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2035
      if "w \<in> ball x r" "w \<noteq> x" for w
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2036
      using that D_eq[of w] n by (auto simp: D_def power_int_diff power_int_minus powr_nat' divide_simps)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2037
  qed (use r n in \<open>auto intro!: holomorphic_intros\<close>)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2038
  thus "zorder f' x = zorder f x - 1"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2039
    using n by (simp add: n_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2040
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2041
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2042
lemma
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2043
  assumes "is_pole f (x :: complex)" "isolated_singularity_at f x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2044
  shows   is_pole_deriv: "is_pole (deriv f) x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2045
    and   zorder_deriv:  "zorder (deriv f) x = zorder f x - 1"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2046
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2047
  from assms(2) obtain r where r: "r > 0" "f analytic_on ball x r - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2048
    by (auto simp: isolated_singularity_at_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2049
  hence holo: "f holomorphic_on ball x r - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2050
    by (subst (asm) analytic_on_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2051
  have *: "x \<in> ball x r" "open (ball x r)" "open (ball x r - {x})"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2052
    using \<open>r > 0\<close> by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2053
  show "is_pole (deriv f) x" "zorder (deriv f) x = zorder f x - 1"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2054
    by (meson "*" assms(1) holo holomorphic_derivI is_pole_deriv' zorder_deriv')+
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2055
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2056
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2057
lemma removable_singularity_deriv':
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2058
  assumes "f \<midarrow>x\<rightarrow> c" "x \<in> A" "open (A :: complex set)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2059
  assumes "\<And>y. y \<in> A - {x} \<Longrightarrow> (f has_field_derivative f' y) (at y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2060
  shows   "\<exists>c. f' \<midarrow>x\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2061
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2062
  have holo: "f holomorphic_on A - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2063
    using assms by (subst holomorphic_on_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2064
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2065
  define g where "g = (\<lambda>y. if y = x then c else f y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2066
  have deriv_g_eq: "deriv g y = f' y" if "y \<in> A - {x}" for y
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2067
  proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2068
    have ev: "eventually (\<lambda>y. y \<in> A - {x}) (nhds y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2069
      using that assms by (intro eventually_nhds_in_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2070
    have "(f has_field_derivative f' y) (at y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2071
      using assms that by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2072
    also have "?this \<longleftrightarrow> (g has_field_derivative f' y) (at y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2073
      by (intro has_field_derivative_cong_ev refl eventually_mono[OF ev]) (auto simp: g_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2074
    finally show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2075
      by (intro DERIV_imp_deriv assms)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2076
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2077
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2078
  have "g holomorphic_on A"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2079
    unfolding g_def using assms assms(1) holo 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2080
    by (intro removable_singularity) auto
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2081
  hence "deriv g holomorphic_on A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2082
    by (intro holomorphic_deriv assms)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2083
  hence "continuous_on A (deriv g)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2084
    by (meson holomorphic_on_imp_continuous_on)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2085
  hence "(deriv g \<longlongrightarrow> deriv g x) (at x within A)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2086
    using assms by (auto simp: continuous_on_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2087
  also have "?this \<longleftrightarrow> (f' \<longlongrightarrow> deriv g x) (at x within A)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2088
    by (intro filterlim_cong refl) (auto simp: eventually_at_filter deriv_g_eq)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2089
  finally have "f' \<midarrow>x\<rightarrow> deriv g x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2090
    using \<open>open A\<close> \<open>x \<in> A\<close> by (meson tendsto_within_open)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2091
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2092
    by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2093
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2094
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2095
lemma removable_singularity_deriv:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2096
  assumes "f \<midarrow>x\<rightarrow> c" "isolated_singularity_at f x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2097
  shows   "\<exists>c. deriv f \<midarrow>x\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2098
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2099
  from assms(2) obtain r where r: "r > 0" "f analytic_on ball x r - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2100
    by (auto simp: isolated_singularity_at_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2101
  hence holo: "f holomorphic_on ball x r - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2102
    using analytic_imp_holomorphic by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2103
  show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2104
    using assms(1)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2105
  proof (rule removable_singularity_deriv')
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2106
    show "x \<in> ball x r" "open (ball x r)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2107
      using \<open>r > 0\<close> by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2108
  qed (auto intro!: holomorphic_derivI[OF holo])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2109
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2110
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2111
lemma not_essential_deriv':
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2112
  assumes "not_essential f x" "x \<in> A" "open A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2113
  assumes "\<And>y. y \<in> A - {x} \<Longrightarrow> (f has_field_derivative f' y) (at y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2114
  shows   "not_essential f' x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2115
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2116
  have holo: "f holomorphic_on A - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2117
    using assms by (subst holomorphic_on_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2118
  from assms consider "is_pole f x" | c where "f \<midarrow>x\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2119
    by (auto simp: not_essential_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2120
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2121
  proof cases
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2122
    case 1
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2123
    thus ?thesis
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2124
      using assms is_pole_deriv' by blast
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2125
  next
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2126
    case (2 c)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2127
    thus ?thesis
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2128
      by (meson assms removable_singularity_deriv' tendsto_imp_not_essential)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2129
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2130
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2131
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2132
lemma not_essential_deriv[singularity_intros]:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2133
  assumes "not_essential f x" "isolated_singularity_at f x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2134
  shows   "not_essential (deriv f) x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2135
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2136
  from assms(2) obtain r where r: "r > 0" "f analytic_on ball x r - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2137
    by (auto simp: isolated_singularity_at_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2138
  hence holo: "f holomorphic_on ball x r - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2139
    by (subst (asm) analytic_on_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2140
  show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2141
    using assms(1)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2142
  proof (rule not_essential_deriv')
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2143
    show "x \<in> ball x r" "open (ball x r)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2144
      using \<open>r > 0\<close> by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2145
  qed (auto intro!: holomorphic_derivI[OF holo])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2146
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2147
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2148
lemma not_essential_frequently_0_imp_tendsto_0:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2149
  fixes f :: "complex \<Rightarrow> complex"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2150
  assumes sing: "isolated_singularity_at f z" "not_essential f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2151
  assumes freq: "frequently (\<lambda>z. f z = 0) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2152
  shows   "f \<midarrow>z\<rightarrow> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2153
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2154
  from freq obtain g :: "nat \<Rightarrow> complex" where g: "filterlim g (at z) at_top" "\<And>n. f (g n) = 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2155
    using frequently_atE by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2156
  have "eventually (\<lambda>x. f (g x) = 0) sequentially"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2157
    using g by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2158
  hence fg: "(\<lambda>x. f (g x)) \<longlonglongrightarrow> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2159
    by (simp add: tendsto_eventually)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2160
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2161
  from assms(2) consider c where "f \<midarrow>z\<rightarrow> c" | "is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2162
    unfolding not_essential_def by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2163
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2164
  proof cases
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2165
    case (1 c)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2166
    have "(\<lambda>x. f (g x)) \<longlonglongrightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2167
      by (rule filterlim_compose[OF 1 g(1)])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2168
    with fg have "c = 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2169
      using LIMSEQ_unique by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2170
    with 1 show ?thesis by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2171
  next
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2172
    case 2
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2173
    have "filterlim (\<lambda>x. f (g x)) at_infinity sequentially"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2174
      using "2" filterlim_compose g(1) is_pole_def by blast
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2175
    with fg have False
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2176
      by (meson not_tendsto_and_filterlim_at_infinity sequentially_bot)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2177
    thus ?thesis ..
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2178
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2179
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2180
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2181
lemma not_essential_frequently_0_imp_eventually_0:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2182
  fixes f :: "complex \<Rightarrow> complex"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2183
  assumes sing: "isolated_singularity_at f z" "not_essential f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2184
  assumes freq: "frequently (\<lambda>z. f z = 0) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2185
  shows   "eventually (\<lambda>z. f z = 0) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2186
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2187
  from sing obtain r where r: "r > 0" and "f analytic_on ball z r - {z}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2188
    by (auto simp: isolated_singularity_at_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2189
  hence holo: "f holomorphic_on ball z r - {z}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2190
    by (subst (asm) analytic_on_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2191
  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2192
    using r by (intro eventually_at_in_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2193
  from freq and this have "frequently (\<lambda>w. f w = 0 \<and> w \<in> ball z r - {z}) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2194
    using frequently_eventually_frequently by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2195
  hence "frequently (\<lambda>w. w \<in> {w\<in>ball z r - {z}. f w = 0}) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2196
    by (simp add: conj_commute)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2197
  hence limpt: "z islimpt {w\<in>ball z r - {z}. f w = 0}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2198
    using islimpt_conv_frequently_at by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2199
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2200
  define g where "g = (\<lambda>w. if w = z then 0 else f w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2201
  have "f \<midarrow>z\<rightarrow> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2202
    by (intro not_essential_frequently_0_imp_tendsto_0 assms)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2203
  hence g_holo: "g holomorphic_on ball z r"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2204
    unfolding g_def by (intro removable_singularity holo) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2205
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2206
  have g_eq_0: "g w = 0" if "w \<in> ball z r" for w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2207
  proof (rule analytic_continuation[where f = g])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2208
    show "open (ball z r)" "connected (ball z r)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2209
      using r by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2210
    show "z islimpt {w\<in>ball z r - {z}. f w = 0}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2211
      by fact
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2212
    show "g w = 0" if "w \<in> {w \<in> ball z r - {z}. f w = 0}" for w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2213
      using that by (auto simp: g_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2214
  qed (use r that g_holo in auto)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2215
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2216
  have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2217
    using r by (intro eventually_at_in_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2218
  thus "eventually (\<lambda>w. f w = 0) (at z)"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2219
    by (metis freq non_zero_neighbour not_eventually not_frequently sing)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2220
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2221
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2222
lemma pole_imp_not_constant:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2223
  fixes f :: "'a :: {perfect_space} \<Rightarrow> _"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2224
  assumes "is_pole f x" "open A" "x \<in> A" "A \<subseteq> insert x B"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2225
  shows   "\<not>f constant_on B"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2226
proof
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2227
  assume *: "f constant_on B"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2228
  then obtain c where c: "\<forall>x\<in>B. f x = c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2229
    by (auto simp: constant_on_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2230
  have "eventually (\<lambda>y. y \<in> A - {x}) (at x)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2231
    using assms by (intro eventually_at_in_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2232
  hence "eventually (\<lambda>y. f y = c) (at x)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2233
    by eventually_elim (use c assms in auto)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2234
  hence **: "f \<midarrow>x\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2235
    by (simp add: tendsto_eventually)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2236
  show False
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2237
    using ** \<open>is_pole f x\<close> at_neq_bot is_pole_def 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2238
          not_tendsto_and_filterlim_at_infinity by blast
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2239
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2240
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2241
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2242
lemma neg_zorder_imp_is_pole:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2243
  assumes iso: "isolated_singularity_at f z" and f_ness: "not_essential f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2244
      and "zorder f z < 0" and fre_nz: "\<exists>\<^sub>F w in at z. f w \<noteq> 0 "
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2245
    shows "is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2246
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2247
  define P where "P = zor_poly f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2248
  define n where "n = zorder f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2249
  have "n<0" unfolding n_def by (simp add: assms(3))
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2250
  define nn where "nn = nat (-n)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2251
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2252
  obtain r where r: "P z \<noteq> 0" "r>0" and r_holo: "P holomorphic_on cball z r" and
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2253
       w_Pn: "(\<forall>w\<in>cball z r - {z}. f w = P w * (w-z) powi n \<and> P w \<noteq> 0)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2254
    using zorder_exist[OF iso f_ness fre_nz,folded P_def n_def] by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2255
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2256
  have "is_pole (\<lambda>w. P w * (w-z) powi n) z"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2257
    unfolding is_pole_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2258
  proof (rule tendsto_mult_filterlim_at_infinity)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2259
    show "P \<midarrow>z\<rightarrow> P z"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2260
      by (metis \<open>r>0\<close> r_holo centre_in_ball continuous_on_interior 
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2261
                holomorphic_on_imp_continuous_on interior_cball isContD)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2262
    show "P z\<noteq>0" by (simp add: \<open>P z \<noteq> 0\<close>)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2263
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2264
    have "LIM x at z. inverse ((x - z) ^ nat (-n)) :> at_infinity"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2265
      apply (subst filterlim_inverse_at_iff[symmetric])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2266
      using \<open>n<0\<close>
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2267
      by (auto intro!:tendsto_eq_intros filterlim_atI
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2268
              simp add: eventually_at_filter)
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2269
    then show "LIM x at z. (x - z) powi n :> at_infinity"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2270
    proof (elim filterlim_mono_eventually)
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2271
      have "inverse ((x - z) ^ nat (-n)) = (x - z) powi n"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2272
        if "x\<noteq>z" for x
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2273
        by (metis \<open>n < 0\<close> linorder_not_le power_int_def power_inverse)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2274
      then show "\<forall>\<^sub>F x in at z. inverse ((x - z) ^ nat (-n))
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2275
                  = (x - z) powi n"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2276
        by (simp add: eventually_at_filter)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2277
    qed auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2278
  qed
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2279
  moreover have "\<forall>\<^sub>F w in at z. f w =  P w * (w-z) powi n"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2280
    unfolding eventually_at_le
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2281
    using w_Pn \<open>r>0\<close> by (force simp add: dist_commute)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2282
  ultimately show ?thesis using is_pole_cong by fast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2283
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2284
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2285
lemma is_pole_divide_zorder:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2286
  fixes f g:: "complex \<Rightarrow> complex" and z::complex
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2287
  assumes f_iso: "isolated_singularity_at f z" and g_iso: "isolated_singularity_at g z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2288
      and f_ness: "not_essential f z" and g_ness: "not_essential g z"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2289
      and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2290
      and z_less: "zorder f z < zorder g z"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2291
    shows "is_pole (\<lambda>z. f z / g z) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2292
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2293
  define fn gn fg where "fn=zorder f z" and "gn=zorder g z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2294
                        and "fg=(\<lambda>w. f w / g w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2295
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2296
  have "isolated_singularity_at fg z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2297
    unfolding fg_def using f_iso g_iso g_ness
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2298
    by (auto intro: singularity_intros)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2299
  moreover have "not_essential fg z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2300
    unfolding fg_def using f_iso g_iso g_ness f_ness
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2301
    by (auto intro: singularity_intros)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2302
  moreover have "zorder fg z < 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2303
  proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2304
    have "zorder fg z = fn - gn"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2305
      using zorder_divide[OF f_iso g_iso f_ness g_ness fg_nconst]
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2306
      by (simp add: fg_def fn_def gn_def) 
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2307
    then show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2308
      using z_less by (simp add: fn_def gn_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2309
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2310
  moreover have "\<exists>\<^sub>F w in at z. fg w \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2311
    using fg_nconst unfolding fg_def by force
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2312
  ultimately show "is_pole fg z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2313
    using neg_zorder_imp_is_pole by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2314
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2315
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2316
lemma isolated_pole_imp_nzero_times:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2317
  assumes f_iso: "isolated_singularity_at f z"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2318
    and "is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2319
  shows "\<exists>\<^sub>Fw in (at z). deriv f w * f w \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2320
proof (rule ccontr)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2321
  assume "\<not> (\<exists>\<^sub>F w in at z.  deriv f w  * f w \<noteq> 0)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2322
  then have "\<forall>\<^sub>F x in at z. deriv f x * f x = 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2323
    unfolding not_frequently by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2324
  moreover have "\<forall>\<^sub>F w in at z. f w \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2325
    using non_zero_neighbour_pole[OF \<open>is_pole f z\<close>] .
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2326
  moreover have "\<forall>\<^sub>F w in at z. deriv f w \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2327
    using is_pole_deriv[OF \<open>is_pole f z\<close> f_iso,THEN non_zero_neighbour_pole]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2328
    .
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2329
  ultimately have "\<forall>\<^sub>F w in at z. False"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2330
    by eventually_elim auto
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2331
  then show False by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2332
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2333
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2334
lemma isolated_pole_imp_neg_zorder:
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2335
  assumes "isolated_singularity_at f z" and "is_pole f z"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2336
  shows "zorder f z < 0"
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2337
  using analytic_imp_holomorphic assms centre_in_ball isolated_singularity_at_def zorder_exist_pole by blast
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2338
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2339
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2340
lemma isolated_singularity_at_deriv[singularity_intros]:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2341
  assumes "isolated_singularity_at f x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2342
  shows "isolated_singularity_at (deriv f) x"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2343
  by (meson analytic_deriv assms isolated_singularity_at_def)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2344
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2345
lemma zorder_deriv_minus_1:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2346
  fixes f g:: "complex \<Rightarrow> complex" and z::complex
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2347
  assumes f_iso: "isolated_singularity_at f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2348
      and f_ness: "not_essential f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2349
      and f_nconst: "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2350
      and f_ord: "zorder f z \<noteq>0"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2351
    shows "zorder (deriv f) z = zorder f z - 1"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2352
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2353
  define P where "P = zor_poly f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2354
  define n where "n = zorder f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2355
  have "n\<noteq>0" unfolding n_def using f_ord by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2356
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2357
  obtain r where "P z \<noteq> 0" "r>0" and P_holo: "P holomorphic_on cball z r"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2358
          and "(\<forall>w\<in>cball z r - {z}. f w
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2359
                            = P w * (w-z) powi n \<and> P w \<noteq> 0)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2360
    using zorder_exist[OF f_iso f_ness f_nconst,folded P_def n_def] by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2361
  from this(4)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2362
  have f_eq: "(\<forall>w\<in>cball z r - {z}. f w
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2363
                            = P w * (w-z) powi n \<and> P w \<noteq> 0)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2364
    using complex_powr_of_int f_ord n_def by presburger
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2365
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2366
  define D where "D = (\<lambda>w. (deriv P w * (w-z) + of_int n * P w)
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2367
                          * (w-z) powi (n - 1))"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2368
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2369
  have deriv_f_eq: "deriv f w = D w" if "w \<in> ball z r - {z}" for w
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2370
  proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2371
    have ev': "eventually (\<lambda>w. w \<in> ball z r - {z}) (nhds w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2372
      using that by (intro eventually_nhds_in_open) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2373
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2374
    define wz where "wz = w - z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2375
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2376
    have "wz \<noteq>0" unfolding wz_def using that by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2377
    moreover have "(P has_field_derivative deriv P w) (at w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2378
      by (meson DiffD1 Elementary_Metric_Spaces.open_ball P_holo
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2379
          ball_subset_cball holomorphic_derivI holomorphic_on_subset that)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2380
    ultimately have "((\<lambda>w. P w * (w-z) powi n) has_field_derivative D w) (at w)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2381
      unfolding D_def using that
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2382
      apply (auto intro!: derivative_eq_intros)
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2383
      by (auto simp: algebra_simps simp flip:power_int_add_1' wz_def)
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2384
    also have "?this \<longleftrightarrow> (f has_field_derivative D w) (at w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2385
      using f_eq
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2386
      by (intro has_field_derivative_cong_ev refl eventually_mono[OF ev']) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2387
    ultimately have "(f has_field_derivative D w) (at w)" by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2388
    moreover have "(f has_field_derivative deriv f w) (at w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2389
      by (metis DERIV_imp_deriv calculation)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2390
    ultimately show ?thesis using DERIV_imp_deriv by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2391
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2392
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2393
  show "zorder (deriv f) z = n - 1"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2394
  proof (rule zorder_eqI)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2395
    show "open (ball z r)" "z \<in> ball z r"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2396
      using \<open>r > 0\<close> by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2397
    define g where "g=(\<lambda>w. (deriv P w * (w-z) + of_int n * P w))"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2398
    show "g holomorphic_on ball z r"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2399
      unfolding g_def using P_holo
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2400
      by (auto intro!:holomorphic_intros)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2401
    show "g z \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2402
      unfolding g_def using \<open>P z \<noteq> 0\<close> \<open>n\<noteq>0\<close> by auto
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2403
    show "deriv f w = (deriv P w * (w-z) + of_int n * P w) * (w-z) powi (n - 1)"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2404
      if "w \<in> ball z r" "w \<noteq> z" for w
77322
9c295f84d55f Replacing z powr of_int i by z powi i and adding new material from the AFP
paulson <lp15@cam.ac.uk>
parents: 77277
diff changeset
  2405
      using D_def deriv_f_eq that by blast
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2406
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2407
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2408
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2409
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2410
lemma deriv_divide_is_pole: \<comment>\<open>Generalises @{thm zorder_deriv}\<close>
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2411
  fixes f g:: "complex \<Rightarrow> complex" and z::complex
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2412
  assumes f_iso: "isolated_singularity_at f z"
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2413
      and f_ness: "not_essential f z" 
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2414
      and fg_nconst: "\<exists>\<^sub>Fw in (at z). deriv f w *  f w \<noteq> 0"
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2415
      and f_ord: "zorder f z \<noteq>0"
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2416
    shows "is_pole (\<lambda>z. deriv f z / f z) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2417
proof (rule neg_zorder_imp_is_pole)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2418
  define ff where "ff=(\<lambda>w. deriv f w / f w)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2419
  show "isolated_singularity_at ff z" 
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2420
    using f_iso f_ness unfolding ff_def
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2421
    by (auto intro: singularity_intros)
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2422
  show "not_essential ff z" 
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2423
    unfolding ff_def using f_ness f_iso by (auto intro: singularity_intros)
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2424
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2425
  have "zorder ff z =  zorder (deriv f) z - zorder f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2426
    unfolding ff_def using f_iso f_ness fg_nconst
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2427
    using isolated_singularity_at_deriv not_essential_deriv zorder_divide by blast
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2428
  moreover have "zorder (deriv f) z = zorder f z - 1"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2429
    using f_iso f_ness f_ord fg_nconst frequently_elim1 zorder_deriv_minus_1 by fastforce
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2430
  ultimately show "zorder ff z < 0" by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2431
    
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2432
  show "\<exists>\<^sub>F w in at z. ff w \<noteq> 0" 
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2433
    unfolding ff_def using fg_nconst by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2434
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2435
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2436
lemma is_pole_deriv_divide_is_pole:
81899
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2437
  fixes f g:: "complex \<Rightarrow> complex" and z::complex
1171ea4a23e4 more tidying
paulson <lp15@cam.ac.uk>
parents: 79945
diff changeset
  2438
  assumes f_iso: "isolated_singularity_at f z"
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2439
      and "is_pole f z" 
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2440
    shows "is_pole (\<lambda>z. deriv f z / f z) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2441
proof (rule deriv_divide_is_pole[OF f_iso])
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2442
  show "not_essential f z" 
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2443
    using \<open>is_pole f z\<close> unfolding not_essential_def by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2444
  show "\<exists>\<^sub>F w in at z. deriv f w * f w \<noteq> 0"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2445
    using assms f_iso isolated_pole_imp_nzero_times by blast
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2446
  show "zorder f z \<noteq> 0"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2447
    using isolated_pole_imp_neg_zorder assms by fastforce
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2448
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2449
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2450
subsection \<open>Isolated zeroes\<close>
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2451
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2452
definition isolated_zero :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> bool" where
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2453
  "isolated_zero f z \<longleftrightarrow> f z = 0 \<and> eventually (\<lambda>z. f z \<noteq> 0) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2454
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2455
lemma isolated_zero_altdef: "isolated_zero f z \<longleftrightarrow> f z = 0 \<and> \<not>z islimpt {z. f z = 0}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2456
  unfolding isolated_zero_def eventually_at_filter eventually_nhds islimpt_def by blast
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2457
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2458
lemma isolated_zero_mult1:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2459
  assumes "isolated_zero f x" "isolated_zero g x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2460
  shows   "isolated_zero (\<lambda>x. f x * g x) x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2461
proof -
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2462
  have "\<forall>\<^sub>F x in at x. f x \<noteq> 0" "\<forall>\<^sub>F x in at x. g x \<noteq> 0"
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2463
    using assms unfolding isolated_zero_def by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2464
  hence "eventually (\<lambda>x. f x * g x \<noteq> 0) (at x)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2465
    by eventually_elim auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2466
  with assms show ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2467
    by (auto simp: isolated_zero_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2468
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2469
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2470
lemma isolated_zero_mult2:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2471
  assumes "isolated_zero f x" "g x \<noteq> 0" "g analytic_on {x}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2472
  shows   "isolated_zero (\<lambda>x. f x * g x) x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2473
proof -
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2474
  have "eventually (\<lambda>x. f x \<noteq> 0) (at x)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2475
    using assms unfolding isolated_zero_def by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2476
  moreover have "eventually (\<lambda>x. g x \<noteq> 0) (at x)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2477
    using analytic_at_neq_imp_eventually_neq[of g x 0] assms by auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2478
  ultimately have "eventually (\<lambda>x. f x * g x \<noteq> 0) (at x)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2479
    by eventually_elim auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2480
  thus ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2481
    using assms(1) by (auto simp: isolated_zero_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2482
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2483
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2484
lemma isolated_zero_mult3:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2485
  assumes "isolated_zero f x" "g x \<noteq> 0" "g analytic_on {x}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2486
  shows   "isolated_zero (\<lambda>x. g x * f x) x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2487
  using isolated_zero_mult2[OF assms] by (simp add: mult_ac)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2488
  
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2489
lemma isolated_zero_prod:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2490
  assumes "\<And>x. x \<in> I \<Longrightarrow> isolated_zero (f x) z" "I \<noteq> {}" "finite I"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2491
  shows   "isolated_zero (\<lambda>y. \<Prod>x\<in>I. f x y) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2492
  using assms(3,2,1) by (induction rule: finite_ne_induct) (auto intro: isolated_zero_mult1)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2493
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2494
lemma non_isolated_zero':
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2495
  assumes "isolated_singularity_at f z" "not_essential f z" "f z = 0" "\<not>isolated_zero f z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2496
  shows   "eventually (\<lambda>z. f z = 0) (at z)"
78517
28c1f4f5335f Numerous minor tweaks and simplifications
paulson <lp15@cam.ac.uk>
parents: 77322
diff changeset
  2497
  by (metis assms isolated_zero_def non_zero_neighbour not_eventually)
77226
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2498
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2499
lemma non_isolated_zero:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2500
  assumes "\<not>isolated_zero f z" "f analytic_on {z}" "f z = 0"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2501
  shows   "eventually (\<lambda>z. f z = 0) (nhds z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2502
proof -
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2503
  have "eventually (\<lambda>z. f z = 0) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2504
    by (rule non_isolated_zero')
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2505
       (use assms in \<open>auto intro: not_essential_analytic isolated_singularity_at_analytic\<close>)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2506
  with \<open>f z = 0\<close> show ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2507
    unfolding eventually_at_filter by (auto elim!: eventually_mono)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2508
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2509
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2510
lemma not_essential_compose:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2511
  assumes "not_essential f (g z)" "g analytic_on {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2512
  shows   "not_essential (\<lambda>x. f (g x)) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2513
proof (cases "isolated_zero (\<lambda>w. g w - g z) z")
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2514
  case False
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2515
  hence "eventually (\<lambda>w. g w - g z = 0) (nhds z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2516
    by (rule non_isolated_zero) (use assms in \<open>auto intro!: analytic_intros\<close>)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2517
  hence "not_essential (\<lambda>x. f (g x)) z \<longleftrightarrow> not_essential (\<lambda>_. f (g z)) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2518
    by (intro not_essential_cong refl)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2519
       (auto elim!: eventually_mono simp: eventually_at_filter)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2520
  thus ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2521
    by (simp add: not_essential_const)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2522
next
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2523
  case True
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2524
  hence ev: "eventually (\<lambda>w. g w \<noteq> g z) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2525
    by (auto simp: isolated_zero_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2526
  from assms consider c where "f \<midarrow>g z\<rightarrow> c" | "is_pole f (g z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2527
    by (auto simp: not_essential_def)  
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2528
  have "isCont g z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2529
    by (rule analytic_at_imp_isCont) fact
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2530
  hence lim: "g \<midarrow>z\<rightarrow> g z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2531
    using isContD by blast
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2532
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2533
  from assms(1) consider c where "f \<midarrow>g z\<rightarrow> c" | "is_pole f (g z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2534
    unfolding not_essential_def by blast
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2535
  thus ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2536
  proof cases
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2537
    fix c assume "f \<midarrow>g z\<rightarrow> c"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2538
    hence "(\<lambda>x. f (g x)) \<midarrow>z\<rightarrow> c"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2539
      by (rule filterlim_compose) (use lim ev in \<open>auto simp: filterlim_at\<close>)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2540
    thus ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2541
      by (auto simp: not_essential_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2542
  next
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2543
    assume "is_pole f (g z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2544
    hence "is_pole (\<lambda>x. f (g x)) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2545
      by (rule is_pole_compose) fact+
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2546
    thus ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2547
      by (auto simp: not_essential_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2548
  qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2549
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2550
  
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2551
subsection \<open>Isolated points\<close>
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2552
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2553
definition isolated_points_of :: "complex set \<Rightarrow> complex set" where
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2554
  "isolated_points_of A = {z\<in>A. eventually (\<lambda>w. w \<notin> A) (at z)}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2555
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2556
lemma isolated_points_of_altdef: "isolated_points_of A = {z\<in>A. \<not>z islimpt A}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2557
  unfolding isolated_points_of_def islimpt_def eventually_at_filter eventually_nhds by blast
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2558
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2559
lemma isolated_points_of_empty [simp]: "isolated_points_of {} = {}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2560
  and isolated_points_of_UNIV [simp]:  "isolated_points_of UNIV = {}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2561
  by (auto simp: isolated_points_of_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2562
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2563
lemma isolated_points_of_open_is_empty [simp]: "open A \<Longrightarrow> isolated_points_of A = {}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2564
  unfolding isolated_points_of_altdef 
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2565
  by (simp add: interior_limit_point interior_open)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2566
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2567
lemma isolated_points_of_subset: "isolated_points_of A \<subseteq> A"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2568
  by (auto simp: isolated_points_of_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2569
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2570
lemma isolated_points_of_discrete:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2571
  assumes "discrete A"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2572
  shows   "isolated_points_of A = A"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2573
  using assms by (auto simp: isolated_points_of_def discrete_altdef)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2574
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2575
lemmas uniform_discreteI1 = uniformI1
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2576
lemmas uniform_discreteI2 = uniformI2
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2577
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2578
lemma isolated_singularity_at_compose:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2579
  assumes "isolated_singularity_at f (g z)" "g analytic_on {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2580
  shows   "isolated_singularity_at (\<lambda>x. f (g x)) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2581
proof (cases "isolated_zero (\<lambda>w. g w - g z) z")
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2582
  case False
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2583
  hence "eventually (\<lambda>w. g w - g z = 0) (nhds z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2584
    by (rule non_isolated_zero) (use assms in \<open>auto intro!: analytic_intros\<close>)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2585
  hence "isolated_singularity_at (\<lambda>x. f (g x)) z \<longleftrightarrow> isolated_singularity_at (\<lambda>_. f (g z)) z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2586
    by (intro isolated_singularity_at_cong refl)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2587
       (auto elim!: eventually_mono simp: eventually_at_filter)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2588
  thus ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2589
    by (simp add: isolated_singularity_at_const)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2590
next
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2591
  case True
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2592
  from assms(1) obtain r where r: "r > 0" "f analytic_on ball (g z) r - {g z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2593
    by (auto simp: isolated_singularity_at_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2594
  hence holo_f: "f holomorphic_on ball (g z) r - {g z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2595
    by (subst (asm) analytic_on_open) auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2596
  from assms(2) obtain r' where r': "r' > 0" "g holomorphic_on ball z r'"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2597
    by (auto simp: analytic_on_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2598
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2599
  have "continuous_on (ball z r') g"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2600
    using holomorphic_on_imp_continuous_on r' by blast
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2601
  hence "isCont g z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2602
    using r' by (subst (asm) continuous_on_eq_continuous_at) auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2603
  hence "g \<midarrow>z\<rightarrow> g z"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2604
    using isContD by blast
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2605
  hence "eventually (\<lambda>w. g w \<in> ball (g z) r) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2606
    using \<open>r > 0\<close> unfolding tendsto_def by force
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2607
  moreover have "eventually (\<lambda>w. g w \<noteq> g z) (at z)" using True
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2608
    by (auto simp: isolated_zero_def elim!: eventually_mono)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2609
  ultimately have "eventually (\<lambda>w. g w \<in> ball (g z) r - {g z}) (at z)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2610
    by eventually_elim auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2611
  then obtain r'' where r'': "r'' > 0" "\<forall>w\<in>ball z r''-{z}. g w \<in> ball (g z) r - {g z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2612
    unfolding eventually_at_filter eventually_nhds_metric ball_def
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2613
    by (auto simp: dist_commute)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2614
  have "f \<circ> g holomorphic_on ball z (min r' r'') - {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2615
  proof (rule holomorphic_on_compose_gen)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2616
    show "g holomorphic_on ball z (min r' r'') - {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2617
      by (rule holomorphic_on_subset[OF r'(2)]) auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2618
    show "f holomorphic_on ball (g z) r - {g z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2619
      by fact
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2620
    show "g ` (ball z (min r' r'') - {z}) \<subseteq> ball (g z) r - {g z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2621
      using r'' by force
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2622
  qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2623
  hence "f \<circ> g analytic_on ball z (min r' r'') - {z}"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2624
    by (subst analytic_on_open) auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2625
  thus ?thesis using \<open>r' > 0\<close> \<open>r'' > 0\<close>
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2626
    by (auto simp: isolated_singularity_at_def o_def intro!: exI[of _ "min r' r''"])
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2627
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2628
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2629
lemma is_pole_power_int_0:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2630
  assumes "f analytic_on {x}" "isolated_zero f x" "n < 0"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2631
  shows   "is_pole (\<lambda>x. f x powi n) x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2632
proof -
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2633
  have "f \<midarrow>x\<rightarrow> f x"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2634
    using assms(1) by (simp add: analytic_at_imp_isCont isContD)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2635
  with assms show ?thesis
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2636
    unfolding is_pole_def
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2637
    by (intro filterlim_power_int_neg_at_infinity) (auto simp: isolated_zero_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2638
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2639
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2640
lemma isolated_zero_imp_not_constant_on:
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2641
  assumes "isolated_zero f x" "x \<in> A" "open A"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2642
  shows   "\<not>f constant_on A"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2643
proof
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2644
  assume "f constant_on A"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2645
  then obtain c where c: "\<And>x. x \<in> A \<Longrightarrow> f x = c"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2646
    by (auto simp: constant_on_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2647
  from assms and c[of x] have [simp]: "c = 0"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2648
    by (auto simp: isolated_zero_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2649
  have "eventually (\<lambda>x. f x \<noteq> 0) (at x)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2650
    using assms by (auto simp: isolated_zero_def)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2651
  moreover have "eventually (\<lambda>x. x \<in> A) (at x)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2652
    using assms by (intro eventually_at_in_open') auto
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2653
  ultimately have "eventually (\<lambda>x. False) (at x)"
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2654
    by eventually_elim (use c in auto)
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2655
  thus False
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2656
    by simp
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2657
qed
69956724ad4f More material for Analysis and Complex_Analysis
paulson <lp15@cam.ac.uk>
parents: 77223
diff changeset
  2658
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2659
end