src/HOL/Complex_Analysis/Complex_Singularities.thy
author paulson <lp15@cam.ac.uk>
Wed, 08 Feb 2023 15:05:24 +0000
changeset 77223 607e1e345e8f
parent 76900 830597d13d6d
child 77226 69956724ad4f
permissions -rw-r--r--
Lots of new material chiefly about complex analysis
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
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     7
definition\<^marker>\<open>tag important\<close> is_pole ::
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     8
  "('a::topological_space \<Rightarrow> 'b::real_normed_vector) \<Rightarrow> 'a \<Rightarrow> bool" where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
     9
  "is_pole f a =  (LIM x (at a). f x :> at_infinity)"
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:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    27
  fixes f::"('a::topological_space \<Rightarrow> 'b::real_normed_div_algebra)"
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
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    30
  by (auto simp add:filterlim_inverse_at_iff[symmetric] comp_def filterlim_at)
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
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    42
lemma is_pole_inverse_holomorphic:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    43
  assumes "open s"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    44
    and f_holo:"f holomorphic_on (s-{z})"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    45
    and pole:"is_pole f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    46
    and non_z:"\<forall>x\<in>s-{z}. f x\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    47
  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
    48
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    49
  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
    50
  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
    51
    by (simp add: g_def cong: LIM_cong)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    52
  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
    53
  hence "continuous_on (s-{z}) (inverse o f)" unfolding comp_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    54
    by (auto elim!:continuous_on_inverse simp add:non_z)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    55
  hence "continuous_on (s-{z}) g" unfolding g_def
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
    56
    using continuous_on_eq by fastforce
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    57
  ultimately have "continuous_on s g" using open_delete[OF \<open>open s\<close>] \<open>open s\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    58
    by (auto simp add:continuous_on_eq_continuous_at)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    59
  moreover have "(inverse o f) holomorphic_on (s-{z})"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    60
    unfolding comp_def using f_holo
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    61
    by (auto elim!:holomorphic_on_inverse simp add:non_z)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    62
  hence "g holomorphic_on (s-{z})"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
    63
    using g_def holomorphic_transform by force
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    64
  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
    65
    by (auto elim!: no_isolated_singularity)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    66
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    67
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    68
lemma not_is_pole_holomorphic:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    69
  assumes "open A" "x \<in> A" "f holomorphic_on A"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    70
  shows   "\<not>is_pole f x"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    71
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    72
  have "continuous_on A f" by (intro holomorphic_on_imp_continuous_on) fact
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    73
  with assms have "isCont f x" by (simp add: continuous_on_eq_continuous_at)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    74
  hence "f \<midarrow>x\<rightarrow> f x" by (simp add: isCont_def)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    75
  thus "\<not>is_pole f x" unfolding is_pole_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    76
    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
    77
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    78
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    79
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
    80
  unfolding is_pole_def inverse_eq_divide [symmetric]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
    81
  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
    82
     (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
    83
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    84
lemma is_pole_cmult_iff [simp]:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    85
  "c \<noteq> 0 \<Longrightarrow> is_pole (\<lambda>z. c * 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
    86
proof
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    87
  assume *: "c \<noteq> 0" "is_pole (\<lambda>z. c * f z) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    88
  have "is_pole (\<lambda>z. inverse c * (c * f z)) z" unfolding is_pole_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    89
    by (rule tendsto_mult_filterlim_at_infinity tendsto_const)+ (use * in \<open>auto simp: is_pole_def\<close>)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    90
  thus "is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    91
    using *(1) by (simp add: field_simps)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    92
next
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    93
  assume *: "c \<noteq> 0" "is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    94
  show "is_pole (\<lambda>z. c * f z) z" unfolding is_pole_def 
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    95
    by (rule tendsto_mult_filterlim_at_infinity tendsto_const)+ (use * in \<open>auto simp: is_pole_def\<close>)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    96
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    97
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
    98
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
    99
  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
   100
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   101
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
   102
  using is_pole_inverse_power[of 1 a] by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   103
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   104
lemma is_pole_divide:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   105
  fixes f :: "'a :: t2_space \<Rightarrow> 'b :: real_normed_field"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   106
  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
   107
  shows   "is_pole (\<lambda>z. f z / g z) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   108
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   109
  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
   110
    using assms filterlim_compose filterlim_inverse_at_infinity isCont_def
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   111
      tendsto_mult_filterlim_at_infinity by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   112
  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
   113
qed
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_basic:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   116
  assumes "f holomorphic_on A" "open A" "z \<in> A" "f z \<noteq> 0" "n > 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   117
  shows   "is_pole (\<lambda>w. f w / (w - z) ^ n) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   118
proof (rule is_pole_divide)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   119
  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
   120
  with assms show "isCont f z" by (auto simp: continuous_on_eq_continuous_at)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   121
  have "filterlim (\<lambda>w. (w - z) ^ n) (nhds 0) (at z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   122
    using assms by (auto intro!: tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   123
  thus "filterlim (\<lambda>w. (w - z) ^ n) (at 0) (at z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   124
    by (intro filterlim_atI tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   125
       (insert assms, auto simp: eventually_at_filter)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   126
qed fact+
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   127
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   128
lemma is_pole_basic':
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   129
  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
   130
  shows   "is_pole (\<lambda>w. f w / w ^ n) 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   131
  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
   132
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   133
text \<open>The proposition
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   134
              \<^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
   135
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
   136
(i.e. the singularity is either removable or a pole).\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   137
definition not_essential::"[complex \<Rightarrow> complex, complex] \<Rightarrow> bool" where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   138
  "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
   139
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   140
definition isolated_singularity_at::"[complex \<Rightarrow> complex, complex] \<Rightarrow> bool" where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   141
  "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
   142
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   143
lemma removable_singularity:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   144
  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
   145
  assumes "f \<midarrow>x\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   146
  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
   147
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   148
  have "continuous_on A ?g"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   149
    unfolding continuous_on_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   150
  proof
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   151
    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
   152
    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
   153
    proof (cases "y = x")
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   154
      case False
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   155
      have "continuous_on (A - {x}) f"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   156
        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
   157
      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
   158
        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
   159
      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
   160
        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
   161
      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
   162
        using y assms False
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   163
        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
   164
      finally show ?thesis .
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   165
    next
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   166
      case [simp]: True
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   167
      have "f \<midarrow>x\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   168
        by fact
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   169
      also have "?this \<longleftrightarrow> (?g \<longlongrightarrow> ?g y) (at y)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   170
        by (intro filterlim_cong) (auto simp: eventually_at_filter)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   171
      finally show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   172
        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
   173
    qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   174
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   175
  moreover {
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   176
    have "?g holomorphic_on A - {x}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   177
      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
   178
  }
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   179
  ultimately show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   180
    by (rule no_isolated_singularity) (use assms in auto)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   181
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   182
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   183
lemma removable_singularity':
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   184
  assumes "isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   185
  assumes "f \<midarrow>z\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   186
  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
   187
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   188
  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
   189
    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
   190
  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
   191
  proof (rule removable_singularity)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   192
    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
   193
      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
   194
  qed (use assms in auto)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   195
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   196
    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
   197
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   198
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   199
named_theorems singularity_intros "introduction rules for singularities"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   200
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   201
lemma holomorphic_factor_unique:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   202
  fixes f::"complex \<Rightarrow> complex" and z::complex and r::real and m n::int
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   203
  assumes "r>0" "g z\<noteq>0" "h z\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   204
    and asm:"\<forall>w\<in>ball z r-{z}. f w = g w * (w-z) powr n \<and> g w\<noteq>0 \<and> f w =  h w * (w - z) powr m \<and> h w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   205
    and g_holo:"g holomorphic_on ball z r" and h_holo:"h holomorphic_on ball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   206
  shows "n=m"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   207
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   208
  have [simp]:"at z within ball z r \<noteq> bot" using \<open>r>0\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   209
      by (auto simp add:at_within_ball_bot_iff)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   210
  have False when "n>m"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   211
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   212
    have "(h \<longlongrightarrow> 0) (at z within ball z r)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   213
    proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) powr (n - m) * g w"])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   214
      have "\<forall>w\<in>ball z r-{z}. h w = (w-z)powr(n-m) * g w"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   215
        using \<open>n>m\<close> asm \<open>r>0\<close> by (simp add: field_simps powr_diff) force
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   216
      then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   217
            \<Longrightarrow> (x' - z) powr (n - m) * g x' = h x'" for x' by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   218
    next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   219
      define F where "F \<equiv> at z within ball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   220
      define f' where "f' \<equiv> \<lambda>x. (x - z) powr (n-m)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   221
      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
   222
      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
   223
        using \<open>n>m\<close>
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   224
          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
   225
      ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   226
        by (simp add: continuous_within)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   227
      moreover have "(g \<longlongrightarrow> g z) F"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   228
        using holomorphic_on_imp_continuous_on[OF g_holo,unfolded continuous_on_def] \<open>r>0\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   229
        unfolding F_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   230
      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
   231
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   232
    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
   233
      using holomorphic_on_imp_continuous_on[OF h_holo]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   234
      by (auto simp add:continuous_on_def \<open>r>0\<close>)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   235
    ultimately have "h z=0" by (auto intro!: tendsto_unique)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   236
    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
   237
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   238
  moreover have False when "m>n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   239
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   240
    have "(g \<longlongrightarrow> 0) (at z within ball z r)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   241
    proof (rule Lim_transform_within[OF _ \<open>r>0\<close>, where f="\<lambda>w. (w - z) powr (m - n) * h w"])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   242
      have "\<forall>w\<in>ball z r -{z}. g w = (w-z) powr (m-n) * h w" using \<open>m>n\<close> asm
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   243
        by (simp add:field_simps powr_diff) force
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   244
      then show "\<lbrakk>x' \<in> ball z r; 0 < dist x' z;dist x' z < r\<rbrakk>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   245
            \<Longrightarrow> (x' - z) powr (m - n) * h x' = g x'" for x' by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   246
    next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   247
      define F where "F \<equiv> at z within ball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   248
      define f' where "f' \<equiv>\<lambda>x. (x - z) powr (m-n)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   249
      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
   250
      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
   251
        using \<open>m>n\<close>
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   252
        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
   253
      ultimately have "(f' \<longlongrightarrow> 0) F" unfolding F_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   254
        by (simp add: continuous_within)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   255
      moreover have "(h \<longlongrightarrow> h z) F"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   256
        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
   257
        unfolding F_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   258
      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
   259
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   260
    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
   261
      using holomorphic_on_imp_continuous_on[OF g_holo]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   262
      by (auto simp add:continuous_on_def \<open>r>0\<close>)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   263
    ultimately have "g z=0" by (auto intro!: tendsto_unique)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   264
    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
   265
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   266
  ultimately show "n=m" by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   267
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   268
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   269
lemma holomorphic_factor_puncture:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   270
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   271
      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>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   272
      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>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   273
  shows "\<exists>!n::int. \<exists>g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   274
          \<and> (\<forall>w\<in>cball z r-{z}. f w = g w * (w-z) powr n \<and> g w\<noteq>0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   275
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   276
  define P where "P = (\<lambda>f n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   277
          \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powr (of_int n)  \<and> g w\<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   278
  have imp_unique:"\<exists>!n::int. \<exists>g r. P f n g r" when "\<exists>n g r. P f n g r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   279
  proof (rule ex_ex1I[OF that])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   280
    fix n1 n2 :: int
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   281
    assume g1_asm:"\<exists>g1 r1. P f n1 g1 r1" and g2_asm:"\<exists>g2 r2. P f n2 g2 r2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   282
    define fac where "fac \<equiv> \<lambda>n g r. \<forall>w\<in>cball z r-{z}. f w = g w * (w - z) powr (of_int n) \<and> g w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   283
    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
   284
        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
   285
    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
   286
        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
   287
    define r where "r \<equiv> min r1 r2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   288
    have "r>0" using \<open>r1>0\<close> \<open>r2>0\<close> unfolding r_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   289
    moreover have "\<forall>w\<in>ball z r-{z}. f w = g1 w * (w-z) powr n1 \<and> g1 w\<noteq>0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   290
        \<and> f w = g2 w * (w - z) powr n2  \<and> g2 w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   291
      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
   292
      by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   293
    ultimately show "n1=n2" using g1_holo g2_holo \<open>g1 z\<noteq>0\<close> \<open>g2 z\<noteq>0\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   294
      apply (elim holomorphic_factor_unique)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   295
      by (auto simp add:r_def)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   296
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   297
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   298
  have P_exist:"\<exists> n g r. P h n g r" when
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   299
      "\<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
   300
    for h
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   301
  proof -
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   302
    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
   303
      unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   304
    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
   305
    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
   306
    have "h' holomorphic_on ball z r"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   307
    proof (rule no_isolated_singularity'[of "{z}"])
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   308
      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
   309
        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
   310
      show "h' holomorphic_on ball z r - {z}"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   311
        using r analytic_imp_holomorphic h'_def holomorphic_transform by fastforce
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   312
    qed auto
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   313
    have ?thesis when "z'=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   314
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   315
      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
   316
      moreover have "\<not> h' constant_on ball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   317
        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
   318
        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
   319
      moreover note \<open>h' holomorphic_on ball z r\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   320
      ultimately obtain g r1 n where "0 < n" "0 < r1" "ball z r1 \<subseteq> ball z r" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   321
          g:"g holomorphic_on ball z r1"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   322
          "\<And>w. w \<in> ball z r1 \<Longrightarrow> h' w = (w - z) ^ n * g w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   323
          "\<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
   324
        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
   325
                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>]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   326
        by (auto simp add:dist_commute)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   327
      define rr where "rr=r1/2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   328
      have "P h' n g rr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   329
        unfolding P_def rr_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   330
        using \<open>n>0\<close> \<open>r1>0\<close> g by (auto simp add:powr_nat)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   331
      then have "P h n g rr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   332
        unfolding h'_def P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   333
      then show ?thesis unfolding P_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   334
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   335
    moreover have ?thesis when "z'\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   336
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   337
      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
   338
      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
   339
      proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   340
        have "isCont h' z" "h' z\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   341
          by (auto simp add: Lim_cong_within \<open>h \<midarrow>z\<rightarrow> z'\<close> \<open>z'\<noteq>0\<close> continuous_at h'_def)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   342
        then obtain r2 where r2:"r2>0" "\<forall>x\<in>ball z r2. h' x\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   343
          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
   344
        define r1 where "r1=min r2 r / 2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   345
        have "0 < r1" "cball z r1 \<subseteq> ball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   346
          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
   347
        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
   348
          using r2 unfolding r1_def by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   349
        ultimately show ?thesis using that by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   350
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   351
      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
   352
      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
   353
      then show ?thesis unfolding P_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   354
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   355
    ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   356
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   357
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   358
  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
   359
    apply (rule_tac imp_unique[unfolded P_def])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   360
    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
   361
  moreover have ?thesis when "is_pole f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   362
  proof (rule imp_unique[unfolded P_def])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   363
    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"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   364
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   365
      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
   366
        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
   367
        by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   368
      then obtain e1 where e1:"e1>0" "\<forall>x\<in>ball z e1-{z}. f x\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   369
        using that eventually_at[of "\<lambda>x. f x\<noteq>0" z UNIV,simplified] by (auto simp add:dist_commute)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   370
      obtain e2 where e2:"e2>0" "f holomorphic_on ball z e2 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   371
        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
   372
      show ?thesis
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   373
        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
   374
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   375
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   376
    define h where "h \<equiv> \<lambda>x. inverse (f x)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   377
    have "\<exists>n g r. P h n g r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   378
    proof -
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   379
      have "(\<lambda>x. inverse (f x)) analytic_on ball z e - {z}"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   380
        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
   381
      moreover have "h \<midarrow>z\<rightarrow> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   382
        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
   383
      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
   384
        using non_zero by (simp add: h_def)
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   385
      ultimately show ?thesis
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   386
        using P_exist[of h] \<open>e > 0\<close>
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   387
        unfolding isolated_singularity_at_def h_def
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   388
        by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   389
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   390
    then obtain n g r
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   391
      where "0 < r" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   392
            g_holo:"g holomorphic_on cball z r" and "g z\<noteq>0" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   393
            g_fac:"(\<forall>w\<in>cball z r-{z}. h w = g w * (w - z) powr of_int n  \<and> g w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   394
      unfolding P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   395
    have "P f (-n) (inverse o g) r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   396
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   397
      have "f w = inverse (g w) * (w - z) powr of_int (- n)" when "w\<in>cball z r - {z}" for w
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   398
        by (metis g_fac h_def inverse_inverse_eq inverse_mult_distrib of_int_minus
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   399
            powr_minus that)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   400
      then show ?thesis
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   401
        unfolding P_def comp_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   402
        using \<open>r>0\<close> g_holo g_fac \<open>g z\<noteq>0\<close> by (auto intro:holomorphic_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   403
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   404
    then show "\<exists>x g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z \<noteq> 0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   405
                  \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int x  \<and> g w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   406
      unfolding P_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   407
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   408
  ultimately show ?thesis using \<open>not_essential f z\<close> unfolding not_essential_def  by presburger
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   409
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   410
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   411
lemma not_essential_transform:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   412
  assumes "not_essential g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   413
  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
   414
  shows "not_essential f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   415
  using assms unfolding not_essential_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   416
  by (simp add: filterlim_cong is_pole_cong)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   417
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   418
lemma isolated_singularity_at_transform:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   419
  assumes "isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   420
  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
   421
  shows "isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   422
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   423
  obtain r1 where "r1>0" and r1:"g analytic_on ball z r1 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   424
    using assms(1) unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   425
  obtain r2 where "r2>0" and r2:" \<forall>x. x \<noteq> z \<and> dist x z < r2 \<longrightarrow> g x = f x"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   426
    using assms(2) unfolding eventually_at by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   427
  define r3 where "r3=min r1 r2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   428
  have "r3>0" unfolding r3_def using \<open>r1>0\<close> \<open>r2>0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   429
  moreover have "f analytic_on ball z r3 - {z}"
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
    have "g holomorphic_on ball z r3 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   432
      using r1 unfolding r3_def by (subst (asm) analytic_on_open,auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   433
    then have "f holomorphic_on ball z r3 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   434
      using r2 unfolding r3_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   435
      by (auto simp add:dist_commute elim!:holomorphic_transform)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   436
    then show ?thesis by (subst analytic_on_open,auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   437
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   438
  ultimately show ?thesis unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   439
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   440
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   441
lemma not_essential_powr[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   442
  assumes "LIM w (at z). f w :> (at x)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   443
  shows "not_essential (\<lambda>w. (f w) powr (of_int n)) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   444
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   445
  define fp where "fp=(\<lambda>w. (f w) powr (of_int n))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   446
  have ?thesis when "n>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   447
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   448
    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
   449
      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
   450
    then have "fp \<midarrow>z\<rightarrow> x ^ nat n" unfolding fp_def
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   451
      by (smt (verit, best) LIM_equal of_int_of_nat power_eq_0_iff powr_nat that zero_less_nat_eq)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   452
    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
   453
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   454
  moreover have ?thesis when "n=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   455
  proof -
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   456
    have "\<forall>\<^sub>F x in at z. fp x = 1"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   457
      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
   458
    then have "fp \<midarrow>z\<rightarrow> 1"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   459
      by (simp add: tendsto_eventually)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   460
    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
   461
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   462
  moreover have ?thesis when "n<0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   463
  proof (cases "x=0")
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   464
    case True
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   465
    have "(\<lambda>x. f x ^ nat (- n)) \<midarrow>z\<rightarrow> 0"
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   466
      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
   467
    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
   468
      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
   469
    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
   470
      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
   471
    then have "LIM w (at z). fp w :> at_infinity"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   472
    proof (elim filterlim_mono_eventually)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   473
      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
   474
        using filterlim_at_within_not_equal[OF assms,of 0] unfolding fp_def
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   475
        by (smt (verit, ccfv_threshold) eventually_mono powr_of_int that)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   476
    qed auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   477
    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
   478
  next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   479
    case False
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   480
    let ?xx= "inverse (x ^ (nat (-n)))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   481
    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
   482
      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
   483
    then have "fp \<midarrow>z\<rightarrow> ?xx"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   484
      by (smt (verit) LIM_cong fp_def inverse_nonzero_iff_nonzero power_eq_0_iff powr_eq_0_iff
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   485
          powr_of_int that zero_less_nat_eq)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   486
    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
   487
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   488
  ultimately show ?thesis by linarith
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   489
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   490
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   491
lemma isolated_singularity_at_powr[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   492
  assumes "isolated_singularity_at f z" "\<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
   493
  shows "isolated_singularity_at (\<lambda>w. (f w) powr (of_int n)) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   494
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   495
  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
   496
    using assms(1) unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   497
  then have r1:"f holomorphic_on ball z r1 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   498
    using analytic_on_open[of "ball z r1-{z}" f] by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   499
  obtain r2 where "r2>0" and r2:"\<forall>w. w \<noteq> z \<and> dist w z < r2 \<longrightarrow> f w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   500
    using assms(2) unfolding eventually_at by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   501
  define r3 where "r3=min r1 r2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   502
  have "(\<lambda>w. (f w) powr of_int n) holomorphic_on ball z r3 - {z}"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   503
    by (intro holomorphic_on_powr_of_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
   504
  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
   505
  ultimately show ?thesis
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   506
    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
   507
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   508
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   509
lemma non_zero_neighbour:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   510
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   511
      and f_ness:"not_essential f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   512
      and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   513
    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
   514
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   515
  obtain fn fp fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   516
          and fr: "fp holomorphic_on cball z fr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   517
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   518
    using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   519
  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
   520
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   521
    have "f w = fp w * (w - z) powr of_int fn" "fp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   522
      using fr(2)[rule_format, of w] using that by (auto simp add:dist_commute)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   523
    moreover have "(w - z) powr of_int fn \<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   524
      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
   525
    ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   526
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   527
  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
   528
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   529
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   530
lemma non_zero_neighbour_pole:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   531
  assumes "is_pole f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   532
  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
   533
  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
   534
  unfolding is_pole_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   535
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   536
lemma non_zero_neighbour_alt:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   537
  assumes holo: "f holomorphic_on S"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   538
      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
   539
    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
   540
proof (cases "f z = 0")
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   541
  case True
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   542
  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
   543
  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
   544
  then show ?thesis
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   545
    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
   546
next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   547
  case False
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   548
  obtain r1 where r1:"r1>0" "\<forall>y. dist z y < r1 \<longrightarrow> f y \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   549
    using continuous_at_avoid[of z f, OF _ False] assms(2,4) continuous_on_eq_continuous_at
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   550
      holo holomorphic_on_imp_continuous_on by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   551
  obtain r2 where r2:"r2>0" "ball z r2 \<subseteq> S"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   552
    using assms openE by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   553
  show ?thesis unfolding eventually_at
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   554
    by (metis (no_types) dist_commute dual_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
   555
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   556
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   557
lemma not_essential_times[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   558
  assumes f_ness:"not_essential f z" and g_ness:"not_essential g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   559
  assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   560
  shows "not_essential (\<lambda>w. f w * g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   561
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   562
  define fg where "fg = (\<lambda>w. f w * g w)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   563
  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
   564
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   565
    have "\<forall>\<^sub>Fw in (at z). fg w=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   566
      using that[unfolded frequently_def, simplified] unfolding fg_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   567
      by (auto elim: eventually_rev_mp)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   568
    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
   569
    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
   570
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   571
  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"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   572
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   573
    obtain fn fp fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   574
          and fr: "fp holomorphic_on cball z fr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   575
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   576
      using holomorphic_factor_puncture[OF f_iso f_ness f_nconst,THEN ex1_implies_ex] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   577
    obtain gn gp gr where [simp]:"gp z \<noteq> 0" and "gr > 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   578
          and gr: "gp holomorphic_on cball z gr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   579
                  "\<forall>w\<in>cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \<and> gp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   580
      using holomorphic_factor_puncture[OF g_iso g_ness g_nconst,THEN ex1_implies_ex] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   581
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   582
    define r1 where "r1=(min fr gr)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   583
    have "r1>0" unfolding r1_def using  \<open>fr>0\<close> \<open>gr>0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   584
    have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   585
      when "w\<in>ball z r1 - {z}" for w
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 "f w = fp w * (w - z) powr of_int fn" "fp w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   588
        using fr(2)[rule_format,of w] that unfolding r1_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   589
      moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   590
        using gr(2)[rule_format, of w] that unfolding r1_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   591
      ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   592
        unfolding fg_def by (auto simp add:powr_add)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   593
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   594
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   595
    have [intro]: "fp \<midarrow>z\<rightarrow>fp z" "gp \<midarrow>z\<rightarrow>gp z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   596
        using fr(1) \<open>fr>0\<close> gr(1) \<open>gr>0\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   597
        by (meson open_ball ball_subset_cball centre_in_ball
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   598
            continuous_on_eq_continuous_at continuous_within holomorphic_on_imp_continuous_on
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   599
            holomorphic_on_subset)+
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   600
    have ?thesis when "fn+gn>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   601
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   602
      have "(\<lambda>w. (fp w * gp w) * (w - z) ^ (nat (fn+gn))) \<midarrow>z\<rightarrow>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   603
        using that by (auto intro!:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   604
      then have "fg \<midarrow>z\<rightarrow> 0"
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   605
        using Lim_transform_within[OF _ \<open>r1>0\<close>]
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   606
        by (smt (verit, ccfv_SIG) DiffI dist_commute dist_nz fg_times mem_ball powr_of_int right_minus_eq
76895
498d8babe716 Further simplifications
paulson <lp15@cam.ac.uk>
parents: 73932
diff changeset
   607
            singletonD that)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   608
      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
   609
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   610
    moreover have ?thesis when "fn+gn=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   611
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   612
      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
   613
        using that by (auto intro!:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   614
      then have "fg \<midarrow>z\<rightarrow> fp z*gp z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   615
        apply (elim Lim_transform_within[OF _ \<open>r1>0\<close>])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   616
        apply (subst fg_times)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   617
        by (auto simp add:dist_commute that)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   618
      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
   619
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   620
    moreover have ?thesis when "fn+gn<0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   621
    proof -
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   622
      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
   623
        using eventually_at_topological that
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   624
        by (force intro!: tendsto_eq_intros filterlim_atI)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   625
      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
   626
        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
   627
      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
   628
        using filterlim_divide_at_infinity by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   629
      then have "is_pole fg z" unfolding is_pole_def
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   630
        apply (elim filterlim_transform_within[OF _ _ \<open>r1>0\<close>])
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   631
        by (simp_all add: dist_commute fg_times of_int_of_nat powr_minus_divide that)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   632
      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
   633
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   634
    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
   635
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   636
  ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   637
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   638
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   639
lemma not_essential_inverse[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   640
  assumes f_ness:"not_essential f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   641
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   642
  shows "not_essential (\<lambda>w. inverse (f w)) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   643
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   644
  define vf where "vf = (\<lambda>w. inverse (f w))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   645
  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
   646
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   647
    have "\<forall>\<^sub>Fw in (at z). f w=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   648
      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
   649
    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
   650
      unfolding vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   651
    from tendsto_cong[OF this] have "vf \<midarrow>z\<rightarrow>0" unfolding vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   652
    then show ?thesis unfolding not_essential_def vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   653
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   654
  moreover have ?thesis when "is_pole f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   655
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   656
    have "vf \<midarrow>z\<rightarrow>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   657
      using that filterlim_at filterlim_inverse_at_iff unfolding is_pole_def vf_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   658
    then show ?thesis unfolding not_essential_def vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   659
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   660
  moreover have ?thesis when "\<exists>x. f\<midarrow>z\<rightarrow>x " and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   661
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   662
    from that obtain fz where fz:"f\<midarrow>z\<rightarrow>fz" by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   663
    have ?thesis when "fz=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   664
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   665
      have "(\<lambda>w. inverse (vf w)) \<midarrow>z\<rightarrow>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   666
        using fz that unfolding vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   667
      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
   668
        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
   669
        unfolding vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   670
      ultimately have "is_pole vf z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   671
        using filterlim_inverse_at_iff[of vf "at z"] unfolding filterlim_at is_pole_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   672
      then show ?thesis unfolding not_essential_def vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   673
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   674
    moreover have ?thesis when "fz\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   675
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   676
      have "vf \<midarrow>z\<rightarrow>inverse fz"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   677
        using fz that unfolding vf_def by (auto intro:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   678
      then show ?thesis unfolding not_essential_def vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   679
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   680
    ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   681
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   682
  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
   683
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   684
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   685
lemma isolated_singularity_at_inverse[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   686
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   687
      and f_ness:"not_essential f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   688
  shows "isolated_singularity_at (\<lambda>w. inverse (f w)) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   689
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   690
  define vf where "vf = (\<lambda>w. inverse (f w))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   691
  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
   692
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   693
    have "\<forall>\<^sub>Fw in (at z). f w=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   694
      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
   695
    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
   696
      unfolding vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   697
    then obtain d1 where "d1>0" and d1:"\<forall>x. x \<noteq> z \<and> dist x z < d1 \<longrightarrow> vf x = 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   698
      unfolding eventually_at by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   699
    then have "vf holomorphic_on ball z d1-{z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   700
      apply (rule_tac holomorphic_transform[of "\<lambda>_. 0"])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   701
      by (auto simp add:dist_commute)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   702
    then have "vf analytic_on ball z d1 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   703
      by (simp add: analytic_on_open open_delete)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   704
    then show ?thesis using \<open>d1>0\<close> unfolding isolated_singularity_at_def vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   705
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   706
  moreover have ?thesis when f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   707
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   708
    have "\<forall>\<^sub>F w in at z. f w \<noteq> 0" 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
   709
    then obtain d1 where d1:"d1>0" "\<forall>x. x \<noteq> z \<and> dist x z < d1 \<longrightarrow> f x \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   710
      unfolding eventually_at by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   711
    obtain d2 where "d2>0" and d2:"f analytic_on ball z d2 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   712
      using f_iso unfolding isolated_singularity_at_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   713
    define d3 where "d3=min d1 d2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   714
    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
   715
    moreover
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   716
    have "f analytic_on ball z d3 - {z}"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   717
      by (smt (verit, best) Diff_iff analytic_on_analytic_at d2 d3_def mem_ball)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   718
      then have "vf analytic_on ball z d3 - {z}"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   719
      unfolding vf_def
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   720
      by (intro analytic_on_inverse; simp add: d1(2) d3_def dist_commute)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   721
    ultimately show ?thesis unfolding isolated_singularity_at_def vf_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   722
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   723
  ultimately show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   724
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   725
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   726
lemma not_essential_divide[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   727
  assumes f_ness:"not_essential f z" and g_ness:"not_essential g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   728
  assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   729
  shows "not_essential (\<lambda>w. f w / g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   730
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   731
  have "not_essential (\<lambda>w. f w * inverse (g w)) z"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   732
    by (simp add: f_iso f_ness g_iso g_ness isolated_singularity_at_inverse not_essential_inverse not_essential_times)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   733
  then show ?thesis by (simp add:field_simps)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   734
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   735
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   736
lemma
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   737
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   738
      and g_iso:"isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   739
    shows isolated_singularity_at_times[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   740
              "isolated_singularity_at (\<lambda>w. f w * g w) z" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   741
          isolated_singularity_at_add[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   742
              "isolated_singularity_at (\<lambda>w. f w + g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   743
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   744
  obtain d1 d2 where "d1>0" "d2>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   745
      and d1:"f analytic_on ball z d1 - {z}" and d2:"g analytic_on ball z d2 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   746
    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
   747
  define d3 where "d3=min d1 d2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   748
  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
   749
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   750
  have fan: "f analytic_on ball z d3 - {z}"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   751
    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
   752
  have gan: "g analytic_on ball z d3 - {z}"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   753
    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
   754
  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
   755
    using analytic_on_mult fan gan by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   756
  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
   757
    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
   758
  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
   759
    using analytic_on_add fan gan by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   760
  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
   761
    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
   762
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   763
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   764
lemma isolated_singularity_at_uminus[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   765
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   766
  shows "isolated_singularity_at (\<lambda>w. - f w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   767
  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
   768
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   769
lemma isolated_singularity_at_id[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   770
     "isolated_singularity_at (\<lambda>w. w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   771
  unfolding isolated_singularity_at_def by (simp add: gt_ex)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   772
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   773
lemma isolated_singularity_at_minus[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   774
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   775
      and g_iso:"isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   776
    shows "isolated_singularity_at (\<lambda>w. f w - g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   777
  using isolated_singularity_at_uminus[THEN isolated_singularity_at_add[OF f_iso,of "\<lambda>w. - g w"]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   778
        ,OF g_iso] by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   779
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   780
lemma isolated_singularity_at_divide[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   781
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   782
      and g_iso:"isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   783
      and g_ness:"not_essential g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   784
    shows "isolated_singularity_at (\<lambda>w. f w / g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   785
  using isolated_singularity_at_inverse[THEN isolated_singularity_at_times[OF f_iso,
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   786
          of "\<lambda>w. inverse (g w)"],OF g_iso g_ness] by (simp add:field_simps)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   787
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   788
lemma isolated_singularity_at_const[singularity_intros]:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   789
    "isolated_singularity_at (\<lambda>w. c) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   790
  unfolding isolated_singularity_at_def by (simp add: gt_ex)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   791
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   792
lemma isolated_singularity_at_holomorphic:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   793
  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
   794
  shows "isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   795
  using assms unfolding isolated_singularity_at_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   796
  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
   797
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   798
lemma isolated_singularity_at_shift:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   799
  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
   800
  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
   801
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   802
  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
   803
    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
   804
  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
   805
    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
   806
       (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
   807
  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
   808
    by (simp add: o_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   809
  thus ?thesis using r
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   810
    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
   811
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   812
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   813
lemma isolated_singularity_at_shift_iff:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   814
  "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
   815
  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
   816
        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
   817
  by (auto simp: algebra_simps)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   818
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   819
lemma isolated_singularity_at_shift_0:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   820
  "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
   821
  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
   822
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   823
lemma not_essential_shift:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   824
  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
   825
  shows   "not_essential f (z + w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   826
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   827
  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
   828
    unfolding not_essential_def by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   829
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   830
  proof cases
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   831
    case (1 c)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   832
    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
   833
      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
   834
    thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   835
      by (auto simp: not_essential_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   836
  next
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   837
    case 2
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   838
    hence "is_pole f (z + w)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   839
      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
   840
    thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   841
      by (auto simp: not_essential_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   842
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   843
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   844
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   845
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
   846
  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
   847
        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
   848
  by (auto simp: algebra_simps)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   849
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   850
lemma not_essential_shift_0:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   851
  "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
   852
  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
   853
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   854
lemma not_essential_holomorphic:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   855
  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
   856
  shows   "not_essential f x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   857
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   858
  have "continuous_on A f"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   859
    using assms holomorphic_on_imp_continuous_on by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   860
  hence "f \<midarrow>x\<rightarrow> f x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   861
    using assms continuous_on_eq_continuous_at isContD by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   862
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   863
    by (auto simp: not_essential_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   864
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   865
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   866
lemma eventually_not_pole:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   867
  assumes "isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   868
  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
   869
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   870
  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
   871
    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
   872
  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
   873
    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
   874
  thus "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
   875
  proof eventually_elim
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   876
    case (elim w)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   877
    with r show ?case
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   878
      using analytic_imp_holomorphic not_is_pole_holomorphic open_delete by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   879
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   880
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   881
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   882
lemma not_islimpt_poles:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   883
  assumes "isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   884
  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
   885
  using eventually_not_pole [OF assms]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   886
  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
   887
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   888
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
   889
  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
   890
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   891
lemma not_essential_const [singularity_intros]: "not_essential (\<lambda>_. c) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   892
  unfolding not_essential_def by (rule exI[of _ c]) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   893
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   894
lemma not_essential_uminus [singularity_intros]:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   895
  assumes f_ness: "not_essential f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   896
  assumes f_iso:"isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   897
  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
   898
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   899
  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
   900
    by (intro assms singularity_intros)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   901
  thus ?thesis by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   902
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   903
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   904
subsubsection \<open>The order of non-essential singularities (i.e. removable singularities or poles)\<close>
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
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
   907
  "zorder f z = (THE n. (\<exists>h r. r>0 \<and> h holomorphic_on cball z r \<and> h z\<noteq>0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   908
                   \<and> (\<forall>w\<in>cball z r - {z}. f w =  h w * (w-z) powr (of_int n)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   909
                   \<and> h w \<noteq>0)))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   910
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   911
definition\<^marker>\<open>tag important\<close> zor_poly
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   912
    ::"[complex \<Rightarrow> complex, complex] \<Rightarrow> complex \<Rightarrow> complex" where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   913
  "zor_poly f z = (SOME h. \<exists>r. r > 0 \<and> h holomorphic_on cball z r \<and> h z \<noteq> 0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   914
                   \<and> (\<forall>w\<in>cball z r - {z}. f w =  h w * (w - z) powr (zorder f z)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   915
                   \<and> h w \<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   916
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   917
lemma zorder_exist:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   918
  fixes f::"complex \<Rightarrow> complex" and z::complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   919
  defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   920
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   921
      and f_ness:"not_essential f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   922
      and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   923
  shows "g z\<noteq>0 \<and> (\<exists>r. r>0 \<and> g holomorphic_on cball z r
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   924
    \<and> (\<forall>w\<in>cball z r - {z}. f w  = g w * (w-z) powr n  \<and> g w \<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   925
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   926
  define P where "P = (\<lambda>n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   927
          \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powr (of_int n) \<and> g w\<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   928
  have "\<exists>!n. \<exists>g r. P n g r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   929
    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
   930
  then have "\<exists>g r. P n g r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   931
    unfolding n_def P_def zorder_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   932
    by (drule_tac theI',argo)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   933
  then have "\<exists>r. P n g r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   934
    unfolding P_def zor_poly_def g_def n_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   935
    by (drule_tac someI_ex,argo)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   936
  then obtain r1 where "P n g r1" by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   937
  then show ?thesis unfolding P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   938
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   939
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   940
lemma zorder_shift:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   941
  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
   942
  unfolding zorder_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   943
  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
   944
  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
   945
  subgoal for x h r
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   946
    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
   947
    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
   948
    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
   949
       apply (simp_all flip: cball_translation)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   950
    apply (simp add: add.commute)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   951
    done
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   952
  subgoal for x h r
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   953
    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
   954
    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
   955
    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
   956
       apply (simp_all add: flip: cball_translation_subtract)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   957
    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
   958
  done
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   959
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   960
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
   961
  by (rule zorder_shift)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
   962
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   963
lemma
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   964
  fixes f::"complex \<Rightarrow> complex" and z::complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   965
  assumes f_iso:"isolated_singularity_at f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   966
      and f_ness:"not_essential f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   967
      and f_nconst:"\<exists>\<^sub>Fw in (at z). f w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   968
    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
   969
      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
   970
                                                = inverse (zor_poly f z w)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   971
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   972
  define vf where "vf = (\<lambda>w. inverse (f w))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   973
  define fn vfn where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   974
    "fn = zorder f z"  and "vfn = zorder vf z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   975
  define fp vfp where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   976
    "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
   977
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   978
  obtain fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   979
          and fr: "fp holomorphic_on cball z fr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   980
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   981
    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
   982
    by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   983
  have fr_inverse: "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
   984
        and fr_nz: "inverse (fp w) \<noteq> 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   985
    when "w\<in>ball z fr - {z}" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   986
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   987
    have "f w = fp w * (w - z) powr of_int fn" "fp w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   988
      using fr(2)[rule_format,of w] that by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   989
    then show "vf w = (inverse (fp w)) * (w - z) powr (of_int (-fn))" "inverse (fp w)\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   990
      unfolding vf_def by (auto simp add:powr_minus)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   991
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   992
  obtain vfr where [simp]:"vfp z \<noteq> 0" and "vfr>0" and vfr:"vfp holomorphic_on cball z vfr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   993
      "(\<forall>w\<in>cball z vfr - {z}. vf w = vfp w * (w - z) powr of_int vfn \<and> vfp w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   994
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   995
    have "isolated_singularity_at vf z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   996
      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
   997
    moreover have "not_essential vf z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
   998
      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
   999
    moreover have "\<exists>\<^sub>F w in at z. vf w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1000
      using f_nconst unfolding vf_def by (auto elim:frequently_elim1)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1001
    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
  1002
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1003
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1004
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1005
  define r1 where "r1 = min fr vfr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1006
  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
  1007
  show "vfn = - fn"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1008
  proof (rule holomorphic_factor_unique)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1009
    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
  1010
      using fr_nz by force
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1011
    then show "\<forall>w\<in>ball z r1 - {z}.
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1012
               vf w = vfp w * (w - z) powr complex_of_int vfn \<and>
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1013
               vfp w \<noteq> 0 \<and> vf w = inverse (fp w) * (w - z) powr complex_of_int (- fn) \<and>
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1014
               inverse (fp w) \<noteq> 0"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1015
      using fr_inverse r1_def vfr(2)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1016
      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
  1017
    show "vfp holomorphic_on ball z r1"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1018
      using r1_def vfr(1) by auto
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1019
    show "(\<lambda>w. inverse (fp w)) holomorphic_on ball z r1"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1020
      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
  1021
  qed (use \<open>r1>0\<close> in auto)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1022
  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
  1023
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1024
    have "w \<in> ball z fr - {z}" "w \<in> cball z vfr - {z}"  "w\<noteq>z" using that unfolding r1_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1025
    from fr_inverse[OF this(1)] fr_nz[OF this(1)] vfr(2)[rule_format,OF this(2)] \<open>vfn = - fn\<close> \<open>w\<noteq>z\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1026
    show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1027
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1028
  then show "\<forall>\<^sub>Fw in (at z). vfp w = inverse (fp w)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1029
    unfolding eventually_at using \<open>r1>0\<close>
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1030
    by (metis DiffI dist_commute mem_ball singletonD)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1031
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1032
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1033
lemma zor_poly_shift:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1034
  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
  1035
    and ness1: "not_essential f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1036
    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
  1037
  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
  1038
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1039
  obtain r1 where "r1>0" "zor_poly f z z \<noteq> 0" and
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1040
      holo1:"zor_poly f z holomorphic_on cball z r1" and
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1041
      rball1:"\<forall>w\<in>cball z r1 - {z}.
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1042
           f w = zor_poly f z w * (w - z) powr of_int (zorder f z) \<and>
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1043
           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
  1044
    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
  1045
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1046
  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
  1047
  have "isolated_singularity_at ff 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1048
    unfolding ff_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1049
    using iso1 isolated_singularity_at_shift_iff[of f 0 z]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1050
    by (simp add:algebra_simps)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1051
  moreover have "not_essential ff 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1052
    unfolding ff_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1053
    using ness1 not_essential_shift_iff[of f 0 z]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1054
    by (simp add:algebra_simps)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1055
  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
  1056
    unfolding ff_def using nzero1
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1057
    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
  1058
        eventually_mono not_frequently)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1059
  ultimately obtain r2 where "r2>0" "zor_poly ff 0 0 \<noteq> 0" and
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1060
      holo2:"zor_poly ff 0 holomorphic_on cball 0 r2" and
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1061
      rball2:"\<forall>w\<in>cball 0 r2 - {0}.
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1062
           ff w = zor_poly ff 0 w * w powr of_int (zorder ff 0) \<and>
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1063
           zor_poly ff 0 w \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1064
    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
  1065
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1066
  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
  1067
  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
  1068
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1069
  have "zor_poly f z w = zor_poly ff 0 (w - z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1070
    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
  1071
  proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1072
    define n::complex where "n= of_int (zorder f z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1073
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1074
    have "f w = zor_poly f z w * (w - z) powr n"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1075
    proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1076
      have "w\<in>cball z r1 - {z}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1077
        using r_def that by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1078
      from rball1[rule_format, OF this]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1079
      show ?thesis unfolding n_def by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1080
    qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1081
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1082
    moreover have "f w = zor_poly ff 0 (w - z) * (w - z) powr n"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1083
    proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1084
      have "w-z\<in>cball 0 r2 - {0}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1085
        using r_def that by (auto simp:dist_complex_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1086
      from rball2[rule_format, OF this]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1087
      have "ff (w - z) = zor_poly ff 0 (w - z) * (w - z)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1088
                            powr of_int (zorder ff 0)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1089
        by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1090
      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
  1091
        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
  1092
      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
  1093
    qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1094
    ultimately have "zor_poly f z w * (w - z) powr n
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1095
                = zor_poly ff 0 (w - z) * (w - z) powr n"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1096
      by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1097
    moreover have "(w - z) powr n \<noteq>0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1098
      using that by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1099
    ultimately show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1100
      apply (subst (asm) mult_cancel_right)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1101
      by (simp add:ff_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1102
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1103
  then have "\<forall>\<^sub>F w in at z. zor_poly f z w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1104
                  = zor_poly ff 0 (w - z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1105
    unfolding eventually_at
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1106
    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
  1107
    using \<open>r>0\<close> by (auto simp:dist_commute)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1108
  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
  1109
    using holo1[THEN holomorphic_on_imp_continuous_on]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1110
    apply (elim continuous_on_interior)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1111
    using \<open>r1>0\<close> by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1112
  moreover have "isCont (\<lambda>w. zor_poly ff 0 (w - z)) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1113
  proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1114
    have "isCont (zor_poly ff 0) 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1115
      using holo2[THEN holomorphic_on_imp_continuous_on]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1116
      apply (elim continuous_on_interior)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1117
      using \<open>r2>0\<close> by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1118
    then show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1119
      unfolding isCont_iff by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1120
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1121
  ultimately show "\<forall>\<^sub>F w in nhds z. zor_poly f z w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1122
                  = zor_poly ff 0 (w - z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1123
    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
  1124
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1125
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1126
lemma
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1127
  fixes f g::"complex \<Rightarrow> complex" and z::complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1128
  assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1129
      and f_ness:"not_essential f z" and g_ness:"not_essential g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1130
      and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1131
  shows zorder_times:"zorder (\<lambda>w. f w * g w) z = zorder f z + zorder g z" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1132
        zor_poly_times:"\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w * g w) z w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1133
                                                  = zor_poly f z w *zor_poly g z w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1134
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1135
  define fg where "fg = (\<lambda>w. f w * g w)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1136
  define fn gn fgn where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1137
    "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
  1138
  define fp gp fgp where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1139
    "fp = zor_poly f z" and "gp = zor_poly g z" and "fgp = zor_poly fg z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1140
  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"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1141
    using fg_nconst by (auto elim!:frequently_elim1)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1142
  obtain fr where [simp]:"fp z \<noteq> 0" and "fr > 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1143
          and fr: "fp holomorphic_on cball z fr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1144
                  "\<forall>w\<in>cball z fr - {z}. f w = fp w * (w - z) powr of_int fn \<and> fp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1145
    using zorder_exist[OF f_iso f_ness f_nconst,folded fp_def fn_def] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1146
  obtain gr where [simp]:"gp z \<noteq> 0" and "gr > 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1147
          and gr: "gp holomorphic_on cball z gr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1148
                  "\<forall>w\<in>cball z gr - {z}. g w = gp w * (w - z) powr of_int gn \<and> gp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1149
    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
  1150
  define r1 where "r1=min fr gr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1151
  have "r1>0" unfolding r1_def using \<open>fr>0\<close> \<open>gr>0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1152
  have fg_times:"fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" and fgp_nz:"fp w*gp w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1153
    when "w\<in>ball z r1 - {z}" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1154
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1155
    have "f w = fp w * (w - z) powr of_int fn" "fp w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1156
      using fr(2)[rule_format,of w] that unfolding r1_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1157
    moreover have "g w = gp w * (w - z) powr of_int gn" "gp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1158
      using gr(2)[rule_format, of w] that unfolding r1_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1159
    ultimately show "fg w = (fp w * gp w) * (w - z) powr (of_int (fn+gn))" "fp w*gp w\<noteq>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1160
      unfolding fg_def by (auto simp add:powr_add)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1161
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1162
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1163
  obtain fgr where [simp]:"fgp z \<noteq> 0" and "fgr > 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1164
          and fgr: "fgp holomorphic_on cball z fgr"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1165
                  "\<forall>w\<in>cball z fgr - {z}. fg w = fgp w * (w - z) powr of_int fgn \<and> fgp w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1166
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1167
    have "fgp z \<noteq> 0 \<and> (\<exists>r>0. fgp holomorphic_on cball z r
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1168
            \<and> (\<forall>w\<in>cball z r - {z}. fg w = fgp w * (w - z) powr of_int fgn \<and> fgp w \<noteq> 0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1169
      apply (rule zorder_exist[of fg z, folded fgn_def fgp_def])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1170
      subgoal unfolding fg_def using isolated_singularity_at_times[OF f_iso g_iso] .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1171
      subgoal unfolding fg_def using not_essential_times[OF f_ness g_ness f_iso g_iso] .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1172
      subgoal unfolding fg_def using fg_nconst .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1173
      done
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1174
    then show ?thesis using that by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1175
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1176
  define r2 where "r2 = min fgr r1"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1177
  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
  1178
  show "fgn = fn + gn "
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1179
    apply (rule holomorphic_factor_unique[of r2 fgp z "\<lambda>w. fp w * gp w" fg])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1180
    subgoal using \<open>r2>0\<close> by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1181
    subgoal by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1182
    subgoal by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1183
    subgoal
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1184
    proof (rule ballI)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1185
      fix w assume "w \<in> ball z r2 - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1186
      then have "w \<in> ball z r1 - {z}" "w \<in> cball z fgr - {z}"  unfolding r2_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1187
      from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1188
      show "fg w = fgp w * (w - z) powr of_int fgn \<and> fgp w \<noteq> 0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1189
              \<and> fg w = fp w * gp w * (w - z) powr of_int (fn + gn) \<and> fp w * gp w \<noteq> 0" by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1190
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1191
    subgoal using fgr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1192
    subgoal using fr(1) gr(1) unfolding r2_def r1_def by (auto intro!:holomorphic_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1193
    done
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1194
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1195
  have "fgp w = fp w *gp w" when "w\<in>ball z r2-{z}" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1196
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1197
    have "w \<in> ball z r1 - {z}" "w \<in> cball z fgr - {z}" "w\<noteq>z" using that  unfolding r2_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1198
    from fg_times[OF this(1)] fgp_nz[OF this(1)] fgr(2)[rule_format,OF this(2)] \<open>fgn = fn + gn\<close> \<open>w\<noteq>z\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1199
    show ?thesis by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1200
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1201
  then show "\<forall>\<^sub>Fw in (at z). fgp w = fp w * gp w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1202
    using \<open>r2>0\<close> unfolding eventually_at by (auto simp add:dist_commute)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1203
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1204
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1205
lemma
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1206
  fixes f g::"complex \<Rightarrow> complex" and z::complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1207
  assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1208
      and f_ness:"not_essential f z" and g_ness:"not_essential g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1209
      and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1210
  shows zorder_divide:"zorder (\<lambda>w. f w / g w) z = zorder f z - zorder g z" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1211
        zor_poly_divide:"\<forall>\<^sub>Fw in (at z). zor_poly (\<lambda>w. f w / g w) z w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1212
                                                  = zor_poly f z w  / zor_poly g z w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1213
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1214
  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"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1215
    using fg_nconst by (auto elim!:frequently_elim1)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1216
  define vg where "vg=(\<lambda>w. inverse (g w))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1217
  have "zorder (\<lambda>w. f w * vg w) z = zorder f z + zorder vg z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1218
    apply (rule zorder_times[OF f_iso _ f_ness,of vg])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1219
    subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1220
    subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1221
    subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1222
    done
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1223
  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
  1224
    using zorder_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1225
    by (auto simp add:field_simps)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1226
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1227
  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"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1228
    apply (rule zor_poly_times[OF f_iso _ f_ness,of vg])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1229
    subgoal unfolding vg_def using isolated_singularity_at_inverse[OF g_iso g_ness] .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1230
    subgoal unfolding vg_def using not_essential_inverse[OF g_ness g_iso] .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1231
    subgoal unfolding vg_def using fg_nconst by (auto elim!:frequently_elim1)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1232
    done
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1233
  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
  1234
    using zor_poly_inverse[OF g_iso g_ness g_nconst,folded vg_def] unfolding vg_def
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1235
    by eventually_elim (auto simp add:field_simps)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1236
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1237
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1238
lemma zorder_exist_zero:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1239
  fixes f::"complex \<Rightarrow> complex" and z::complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1240
  defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1241
  assumes  holo: "f holomorphic_on s" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1242
          "open s" "connected s" "z\<in>s"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1243
      and non_const: "\<exists>w\<in>s. f w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1244
  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
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1245
    \<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
  1246
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1247
  obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> s" "g holomorphic_on cball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1248
            "(\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1249
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1250
    have "g z \<noteq> 0 \<and> (\<exists>r>0. g holomorphic_on cball z r
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1251
            \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1252
    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
  1253
      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
  1254
        using holo assms(4,6)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1255
        by (meson Diff_subset open_ball analytic_on_holomorphic holomorphic_on_subset openE)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1256
      show "not_essential f z" unfolding not_essential_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1257
        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
  1258
        by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1259
      have "\<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
  1260
      proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1261
        obtain w where "w\<in>s" "f w\<noteq>0" using non_const by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1262
        then show ?thesis
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1263
          by (rule non_zero_neighbour_alt[OF holo \<open>open s\<close> \<open>connected s\<close> \<open>z\<in>s\<close>])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1264
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1265
      then show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1266
        apply (elim eventually_frequentlyE)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1267
        by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1268
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1269
    then obtain r1 where "g z \<noteq> 0" "r1>0" and r1:"g holomorphic_on cball z r1"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1270
            "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1271
      by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1272
    obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> s"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1273
      using assms(4,6) open_contains_cball_eq by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1274
    define r3 where "r3=min r1 r2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1275
    have "r3>0" "cball z r3 \<subseteq> s" using \<open>r1>0\<close> r2 unfolding r3_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1276
    moreover have "g holomorphic_on cball z r3"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1277
      using r1(1) unfolding r3_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1278
    moreover have "(\<forall>w\<in>cball z r3 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1279
      using r1(2) unfolding r3_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1280
    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
  1281
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1282
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1283
  have fz_lim: "f\<midarrow> z \<rightarrow> f z"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1284
    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
  1285
  have gz_lim: "g \<midarrow>z\<rightarrow>g z"
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1286
    by (metis r open_ball at_within_open ball_subset_cball centre_in_ball
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1287
        continuous_on_def continuous_on_subset holomorphic_on_imp_continuous_on)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1288
  have if_0:"if f z=0 then n > 0 else n=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1289
  proof -
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1290
    have "(\<lambda>w. g w * (w - z) powr of_int n) \<midarrow>z\<rightarrow> f z"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1291
      using fz_lim Lim_transform_within_open[where s="ball z r"] r by fastforce
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1292
    then have "(\<lambda>w. (g w * (w - z) powr of_int n) / g w) \<midarrow>z\<rightarrow> f z/g z"
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1293
      using gz_lim \<open>g z \<noteq> 0\<close> tendsto_divide by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1294
    then have powr_tendsto:"(\<lambda>w. (w - z) powr of_int n) \<midarrow>z\<rightarrow> f z/g z"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1295
      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
  1296
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1297
    have ?thesis when "n\<ge>0" "f z=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1298
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1299
      have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> f z/g z"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1300
        using powr_tendsto Lim_transform_within[where d=r]
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1301
        by (fastforce simp: powr_of_int \<open>n\<ge>0\<close> \<open>r>0\<close>)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1302
      then have *:"(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> 0" using \<open>f z=0\<close> by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1303
      moreover have False when "n=0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1304
      proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1305
        have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> 1"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1306
          using \<open>n=0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1307
        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
  1308
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1309
      ultimately show ?thesis using that by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1310
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1311
    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
  1312
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1313
      have False when "n>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1314
      proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1315
        have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> f z/g z"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1316
          using powr_tendsto Lim_transform_within[where d=r]
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1317
          by (fastforce simp add: powr_of_int \<open>n\<ge>0\<close> \<open>r>0\<close>)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1318
        moreover have "(\<lambda>w. (w - z) ^ nat n) \<midarrow>z\<rightarrow> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1319
          using \<open>n>0\<close> by (auto intro!:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1320
        ultimately show False using \<open>f z\<noteq>0\<close> \<open>g z\<noteq>0\<close> using LIM_unique divide_eq_0_iff by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1321
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1322
      then show ?thesis using that by force
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1323
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1324
    moreover have False when "n<0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1325
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1326
      have "(\<lambda>w. inverse ((w - z) ^ nat (- n))) \<midarrow>z\<rightarrow> f z/g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1327
           "(\<lambda>w.((w - z) ^ nat (- n))) \<midarrow>z\<rightarrow> 0"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1328
        apply (smt (verit, ccfv_SIG) LIM_cong eq_iff_diff_eq_0 powr_of_int powr_tendsto that)
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1329
        using that by (auto intro!:tendsto_eq_intros)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1330
      from tendsto_mult[OF this,simplified]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1331
      have "(\<lambda>x. inverse ((x - z) ^ nat (- n)) * (x - z) ^ nat (- n)) \<midarrow>z\<rightarrow> 0" .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1332
      then have "(\<lambda>x. 1::complex) \<midarrow>z\<rightarrow> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1333
        by (elim Lim_transform_within_open[where s=UNIV],auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1334
      then show False using LIM_const_eq by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1335
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1336
    ultimately show ?thesis by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1337
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1338
  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
  1339
  proof (cases "w=z")
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1340
    case True
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1341
    then have "f \<midarrow>z\<rightarrow>f w"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1342
      using fz_lim by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1343
    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
  1344
    proof (elim Lim_transform_within[OF _ \<open>r>0\<close>])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1345
      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
  1346
      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
  1347
        unfolding cball_def by (auto simp add: dist_commute)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1348
      then have "f x = g x * (x - z) powr of_int n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1349
        using r(4)[rule_format,of x] by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1350
      also have "... = g x * (x - z) ^ nat n"
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1351
        by (smt (verit) \<open>x \<noteq> z\<close> if_0 powr_of_int right_minus_eq)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1352
      finally show "f x = g x * (x - z) ^ nat n" .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1353
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1354
    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
  1355
      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
  1356
    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
  1357
    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
  1358
  next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1359
    case False
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1360
    then have "f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1361
      using r(4) that by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1362
    then show ?thesis using False if_0 powr_of_int by (auto split:if_splits)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1363
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1364
  ultimately show ?thesis using r by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1365
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1366
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1367
lemma zorder_exist_pole:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1368
  fixes f::"complex \<Rightarrow> complex" and z::complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1369
  defines "n\<equiv>zorder f z" and "g\<equiv>zor_poly f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1370
  assumes  holo: "f holomorphic_on s-{z}" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1371
          "open s" "z\<in>s"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1372
      and "is_pole f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1373
  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
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1374
    \<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
  1375
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1376
  obtain r where "g z \<noteq> 0" and r: "r>0" "cball z r \<subseteq> s" "g holomorphic_on cball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1377
            "(\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1378
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1379
    have "g z \<noteq> 0 \<and> (\<exists>r>0. g holomorphic_on cball z r
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1380
            \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1381
    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
  1382
      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
  1383
        using holo assms(4,5)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1384
        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
  1385
      show "not_essential f z" unfolding not_essential_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1386
        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
  1387
        by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1388
      from non_zero_neighbour_pole[OF \<open>is_pole f z\<close>] show "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1389
        apply (elim eventually_frequentlyE)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1390
        by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1391
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1392
    then obtain r1 where "g z \<noteq> 0" "r1>0" and r1:"g holomorphic_on cball z r1"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1393
            "(\<forall>w\<in>cball z r1 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1394
      by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1395
    obtain r2 where r2: "r2>0" "cball z r2 \<subseteq> s"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1396
      using assms(4,5) open_contains_cball_eq by metis
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1397
    define r3 where "r3=min r1 r2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1398
    have "r3>0" "cball z r3 \<subseteq> s" using \<open>r1>0\<close> r2 unfolding r3_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1399
    moreover have "g holomorphic_on cball z r3"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1400
      using r1(1) unfolding r3_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1401
    moreover have "(\<forall>w\<in>cball z r3 - {z}. f w = g w * (w - z) powr of_int n \<and> g w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1402
      using r1(2) unfolding r3_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1403
    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
  1404
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1405
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1406
  have "n<0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1407
  proof (rule ccontr)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1408
    assume " \<not> n < 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1409
    define c where "c=(if n=0 then g z else 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1410
    have [simp]:"g \<midarrow>z\<rightarrow> g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1411
      by (metis open_ball at_within_open ball_subset_cball centre_in_ball
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1412
            continuous_on holomorphic_on_imp_continuous_on holomorphic_on_subset r(1) r(3) )
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1413
    have "\<forall>\<^sub>F x in at z. f x = g x * (x - z) ^ nat n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1414
      unfolding eventually_at_topological
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1415
      apply (rule_tac exI[where x="ball z r"])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1416
      using r powr_of_int \<open>\<not> n < 0\<close> by auto
76897
a56e27f98763 continued proof simplification
paulson <lp15@cam.ac.uk>
parents: 76895
diff changeset
  1417
    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
  1418
    proof (cases "n=0")
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1419
      case True
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1420
      then show ?thesis unfolding c_def by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1421
    next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1422
      case False
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1423
      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
  1424
        by (auto intro!:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1425
      from tendsto_mult[OF _ this,of g "g z",simplified]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1426
      show ?thesis unfolding c_def using False by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1427
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1428
    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
  1429
    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
  1430
      unfolding is_pole_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1431
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1432
  moreover have "\<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
  1433
    using r(4) \<open>n<0\<close> powr_of_int
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1434
    by (metis Diff_iff divide_inverse eq_iff_diff_eq_0 insert_iff linorder_not_le)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1435
  ultimately show ?thesis using r(1-3) \<open>g z\<noteq>0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1436
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1437
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1438
lemma zorder_eqI:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1439
  assumes "open s" "z \<in> s" "g holomorphic_on s" "g z \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1440
  assumes fg_eq:"\<And>w. \<lbrakk>w \<in> s;w\<noteq>z\<rbrakk> \<Longrightarrow> f w = g w * (w - z) powr n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1441
  shows   "zorder f z = n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1442
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1443
  have "continuous_on s g" by (rule holomorphic_on_imp_continuous_on) fact
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1444
  moreover have "open (-{0::complex})" by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1445
  ultimately have "open ((g -` (-{0})) \<inter> s)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1446
    unfolding continuous_on_open_vimage[OF \<open>open s\<close>] by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1447
  moreover from assms have "z \<in> (g -` (-{0})) \<inter> s" by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1448
  ultimately obtain r where r: "r > 0" "cball z r \<subseteq>  s \<inter> (g -` (-{0}))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1449
    unfolding open_contains_cball by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1450
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1451
  let ?gg= "(\<lambda>w. g w * (w - z) powr n)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1452
  define P where "P = (\<lambda>n g r. 0 < r \<and> g holomorphic_on cball z r \<and> g z\<noteq>0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1453
          \<and> (\<forall>w\<in>cball z r - {z}. f w = g w * (w-z) powr (of_int n) \<and> g w\<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1454
  have "P n g r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1455
    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
  1456
  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
  1457
  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
  1458
  proof (rule holomorphic_factor_puncture)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1459
    have "ball z r-{z} \<subseteq> s" using r using ball_subset_cball by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1460
    then have "?gg holomorphic_on ball z r-{z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1461
      using \<open>g holomorphic_on s\<close> r by (auto intro!: holomorphic_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1462
    then have "f holomorphic_on ball z r - {z}"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1463
      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
  1464
    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
  1465
      using analytic_on_open open_delete r(1) by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1466
  next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1467
    have "not_essential ?gg z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1468
    proof (intro singularity_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1469
      show "not_essential g z"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1470
        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
  1471
            isCont_def not_essential_def)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1472
      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
  1473
      then show "LIM w at z. w - z :> at 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1474
        unfolding filterlim_at by (auto intro:tendsto_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1475
      show "isolated_singularity_at g z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1476
        by (meson Diff_subset open_ball analytic_on_holomorphic
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1477
            assms holomorphic_on_subset isolated_singularity_at_def openE)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1478
    qed
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1479
    moreover
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1480
    have "\<forall>\<^sub>F w in at z. g w * (w - z) powr n = f w"
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1481
      unfolding eventually_at_topological using assms fg_eq by force
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1482
    ultimately show "not_essential f z"
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1483
      using not_essential_transform by blast
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1484
    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
  1485
    proof (intro strip)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1486
      fix d::real assume "0 < d"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1487
      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
  1488
      have "z' \<noteq> z" " dist z' z < d "
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1489
        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
  1490
      moreover have "f z' \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1491
      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
  1492
        have "z' \<in> cball z r"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1493
          unfolding z'_def using \<open>r>0\<close> \<open>d>0\<close> by (auto simp add:dist_norm)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1494
        then show " z' \<in> s" using r(2) by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1495
        show "g z' * (z' - z) powr of_int n \<noteq> 0"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1496
          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
  1497
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1498
      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
  1499
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1500
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1501
  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
  1502
    by (rule_tac the1_equality)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1503
  then show ?thesis unfolding zorder_def P_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1504
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1505
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1506
lemma simple_zeroI:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1507
  assumes "open s" "z \<in> s" "g holomorphic_on s" "g z \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1508
  assumes "\<And>w. w \<in> s \<Longrightarrow> f w = g w * (w - z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1509
  shows   "zorder f z = 1"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1510
  using assms zorder_eqI by force
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1511
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1512
lemma higher_deriv_power:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1513
  shows   "(deriv ^^ j) (\<lambda>w. (w - z) ^ n) w =
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1514
             pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1515
proof (induction j arbitrary: w)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1516
  case 0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1517
  thus ?case by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1518
next
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1519
  case (Suc j w)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1520
  have "(deriv ^^ Suc j) (\<lambda>w. (w - z) ^ n) w = deriv ((deriv ^^ j) (\<lambda>w. (w - z) ^ n)) w"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1521
    by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1522
  also have "(deriv ^^ j) (\<lambda>w. (w - z) ^ n) =
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1523
               (\<lambda>w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1524
    using Suc by (intro Suc.IH ext)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1525
  also {
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1526
    have "(\<dots> has_field_derivative of_nat (n - j) *
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1527
               pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - Suc j)) (at w)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1528
      using Suc.prems by (auto intro!: derivative_eq_intros)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1529
    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
  1530
                 pochhammer (of_nat (Suc n - Suc j)) (Suc j)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1531
      by (cases "Suc j \<le> n", subst pochhammer_rec)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1532
         (insert Suc.prems, simp_all add: algebra_simps Suc_diff_le pochhammer_0_left)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1533
    finally have "deriv (\<lambda>w. pochhammer (of_nat (Suc n - j)) j * (w - z) ^ (n - j)) w =
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1534
                    \<dots> * (w - z) ^ (n - Suc j)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1535
      by (rule DERIV_imp_deriv)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1536
  }
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1537
  finally show ?case .
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
lemma zorder_zero_eqI:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1541
  assumes  f_holo:"f holomorphic_on s" and "open s" "z \<in> s"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1542
  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
  1543
  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
  1544
  shows   "zorder f z = n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1545
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1546
  obtain r where [simp]:"r>0" and "ball z r \<subseteq> s"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1547
    using \<open>open s\<close> \<open>z\<in>s\<close> openE by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1548
  have nz':"\<exists>w\<in>ball z r. f w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1549
  proof (rule ccontr)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1550
    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
  1551
    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
  1552
      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
  1553
    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
  1554
      by (intro higher_deriv_cong_ev) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1555
    also have "(deriv ^^ nat n) (\<lambda>_. 0) z = 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1556
      by (induction n) simp_all
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1557
    finally show False using nz by contradiction
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1558
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1559
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1560
  define zn g where "zn = zorder f z" and "g = zor_poly f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1561
  obtain e where e_if:"if f z = 0 then 0 < zn else zn = 0" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1562
            [simp]:"e>0" and "cball z e \<subseteq> ball z r" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1563
            g_holo:"g holomorphic_on cball z e" and
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1564
            e_fac:"(\<forall>w\<in>cball z e. f w = g w * (w - z) ^ nat zn \<and> g w \<noteq> 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1565
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1566
    have "f holomorphic_on ball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1567
      using f_holo \<open>ball z r \<subseteq> s\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1568
    from that zorder_exist_zero[of f "ball z r" z,simplified,OF this nz',folded zn_def g_def]
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1569
    show ?thesis by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1570
  qed
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1571
  then obtain "zn \<ge> 0" "g z\<noteq>0"
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1572
    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
  1573
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1574
  define A where "A \<equiv> (\<lambda>i. of_nat (i choose (nat zn)) * fact (nat zn) * (deriv ^^ (i - nat zn)) g z)"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1575
  have deriv_A:"(deriv ^^ i) f z = (if zn \<le> int i then A i else 0)" for i
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1576
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1577
    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
  1578
      using \<open>cball z e \<subseteq> ball z r\<close> \<open>e>0\<close> by (intro eventually_nhds_in_open) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1579
    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
  1580
      using e_fac eventually_mono by fastforce
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1581
    hence "(deriv ^^ i) f z = (deriv ^^ i) (\<lambda>w. (w - z) ^ nat zn * g w) z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1582
      by (intro higher_deriv_cong_ev) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1583
    also have "\<dots> = (\<Sum>j=0..i. of_nat (i choose j) *
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1584
                       (deriv ^^ j) (\<lambda>w. (w - z) ^ nat zn) z * (deriv ^^ (i - j)) g z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1585
      using g_holo \<open>e>0\<close>
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1586
      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
  1587
    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
  1588
                    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
  1589
    proof (intro sum.cong refl, goal_cases)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1590
      case (1 j)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1591
      have "(deriv ^^ j) (\<lambda>w. (w - z) ^ nat zn) z =
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1592
              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
  1593
        by (subst higher_deriv_power) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1594
      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
  1595
        by (auto simp: not_less pochhammer_0_left pochhammer_fact)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1596
      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
  1597
                   (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
  1598
                        * (deriv ^^ (i - nat zn)) g z else 0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1599
        by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1600
      finally show ?case .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1601
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1602
    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
  1603
      by (auto simp: A_def)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1604
    finally show "(deriv ^^ i) f z = \<dots>" .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1605
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1606
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1607
  have False when "n<zn"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1608
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1609
    have "(deriv ^^ nat n) f z = 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1610
      using deriv_A[of "nat n"] that \<open>n\<ge>0\<close> by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1611
    with nz show False by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1612
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1613
  moreover have "n\<le>zn"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1614
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1615
    have "g z \<noteq> 0" using e_fac[rule_format,of z] \<open>e>0\<close> by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1616
    then have "(deriv ^^ nat zn) f z \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1617
      using deriv_A[of "nat zn"] by(auto simp add:A_def)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1618
    then have "nat zn \<ge> nat n" using zero[of "nat zn"] by linarith
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1619
    moreover have "zn\<ge>0" using e_if by (auto split:if_splits)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1620
    ultimately show ?thesis using nat_le_eq_zle by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1621
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1622
  ultimately show ?thesis unfolding zn_def by fastforce
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1623
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1624
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1625
lemma
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1626
  assumes "eventually (\<lambda>z. f z = g z) (at z)" "z = z'"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1627
  shows zorder_cong:"zorder f z = zorder g z'" and zor_poly_cong:"zor_poly f z = zor_poly g z'"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1628
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1629
  define P where "P = (\<lambda>ff n h r. 0 < r \<and> h holomorphic_on cball z r \<and> h z\<noteq>0
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1630
                    \<and> (\<forall>w\<in>cball z r - {z}. ff w = h w * (w-z) powr (of_int n) \<and> h w\<noteq>0))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1631
  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
  1632
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1633
    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
  1634
    proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1635
      from that(1) obtain r1 where r1_P:"P f n h r1" by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1636
      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"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1637
        unfolding eventually_at_le by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1638
      define r where "r=min r1 r2"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1639
      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
  1640
      moreover have "h holomorphic_on cball z r"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1641
        using r1_P unfolding P_def r_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1642
      moreover have "g w = h w * (w - z) powr of_int n \<and> h w \<noteq> 0" when "w\<in>cball z r - {z}" for w
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1643
      proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1644
        have "f w = h w * (w - z) powr of_int n \<and> h w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1645
          using r1_P that unfolding P_def r_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1646
        moreover have "f w=g w" using r2_dist[rule_format,of w] that unfolding r_def
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1647
          by (simp add: dist_commute)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1648
        ultimately show ?thesis by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1649
      qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1650
      ultimately show ?thesis unfolding P_def by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1651
    qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1652
    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
  1653
      by (simp add: eq_commute)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1654
    show ?thesis
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1655
      by (rule iffI[OF *[OF _ assms(1)] *[OF _ eq']])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1656
  qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1657
  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
  1658
      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
  1659
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1660
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1661
lemma zorder_nonzero_div_power:
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1662
  assumes sz: "open s" "z \<in> s" "f holomorphic_on s" "f z \<noteq> 0" and "n > 0"
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1663
  shows  "zorder (\<lambda>w. f w / (w - z) ^ n) z = - n"
76900
830597d13d6d final tidying of theorems
paulson <lp15@cam.ac.uk>
parents: 76897
diff changeset
  1664
  using zorder_eqI [OF sz] by (simp add: powr_minus_divide)
71201
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1665
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1666
lemma zor_poly_eq:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1667
  assumes "isolated_singularity_at f z" "not_essential f z" "\<exists>\<^sub>F w in at z. f w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1668
  shows "eventually (\<lambda>w. zor_poly f z w = f w * (w - z) powr - zorder f z) (at z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1669
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1670
  obtain r where r:"r>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1671
       "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w * (w - z) powr of_int (zorder f z))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1672
    using zorder_exist[OF assms] by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1673
  then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w * (w - z) powr - zorder f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1674
    by (auto simp: field_simps powr_minus)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1675
  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
  1676
    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
  1677
  thus ?thesis by eventually_elim (insert *, auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1678
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1679
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1680
lemma zor_poly_zero_eq:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1681
  assumes "f holomorphic_on s" "open s" "connected s" "z \<in> s" "\<exists>w\<in>s. f w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1682
  shows "eventually (\<lambda>w. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)) (at z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1683
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1684
  obtain r where r:"r>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1685
       "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w * (w - z) ^ nat (zorder f z))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1686
    using zorder_exist_zero[OF assms] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1687
  then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w / (w - z) ^ nat (zorder f z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1688
    by (auto simp: field_simps powr_minus)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1689
  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
  1690
    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
  1691
  thus ?thesis by eventually_elim (insert *, auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1692
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1693
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1694
lemma zor_poly_pole_eq:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1695
  assumes f_iso:"isolated_singularity_at f z" "is_pole f z"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1696
  shows "eventually (\<lambda>w. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)) (at z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1697
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1698
  obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1699
    using f_iso analytic_imp_holomorphic unfolding isolated_singularity_at_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1700
  obtain r where r:"r>0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1701
       "(\<forall>w\<in>cball z r - {z}. f w = zor_poly f z w / (w - z) ^ nat (- zorder f z))"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1702
    using zorder_exist_pole[OF f_holo,simplified,OF \<open>is_pole f z\<close>] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1703
  then have *: "\<forall>w\<in>ball z r - {z}. zor_poly f z w = f w * (w - z) ^ nat (- zorder f z)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1704
    by (auto simp: field_simps)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1705
  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
  1706
    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
  1707
  thus ?thesis by eventually_elim (insert *, auto)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1708
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1709
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1710
lemma zor_poly_eqI:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1711
  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1712
  defines "n \<equiv> zorder f z0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1713
  assumes "isolated_singularity_at f z0" "not_essential f z0" "\<exists>\<^sub>F w in at z0. f w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1714
  assumes lim: "((\<lambda>x. f (g x) * (g x - z0) powr - n) \<longlongrightarrow> c) F"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1715
  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
  1716
  shows   "zor_poly f z0 z0 = c"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1717
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1718
  from zorder_exist[OF assms(2-4)] obtain r where
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1719
    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
  1720
       "\<And>w. w \<in> cball z0 r - {z0} \<Longrightarrow> f w = zor_poly f z0 w * (w - z0) powr n"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1721
    unfolding n_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1722
  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
  1723
    using eventually_at_ball'[of r z0 UNIV] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1724
  hence "eventually (\<lambda>w. zor_poly f z0 w = f w * (w - z0) powr - n) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1725
    by eventually_elim (insert r, auto simp: field_simps powr_minus)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1726
  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
  1727
    using r by (intro holomorphic_on_imp_continuous_on) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1728
  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
  1729
    by (auto simp: continuous_on_eq_continuous_at)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1730
  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
  1731
    unfolding isCont_def .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1732
  ultimately have "((\<lambda>w. f w * (w - z0) powr - n) \<longlongrightarrow> zor_poly f z0 z0) (at z0)"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1733
    by (blast intro: Lim_transform_eventually)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1734
  hence "((\<lambda>x. f (g x) * (g x - z0) powr - n) \<longlongrightarrow> zor_poly f z0 z0) F"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1735
    by (rule filterlim_compose[OF _ g])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1736
  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
  1737
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1738
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1739
lemma zor_poly_zero_eqI:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1740
  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1741
  defines "n \<equiv> zorder f z0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1742
  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
  1743
  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
  1744
  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
  1745
  shows   "zor_poly f z0 z0 = c"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1746
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1747
  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
  1748
    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
  1749
       "\<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
  1750
    unfolding n_def by blast
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1751
  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
  1752
    using eventually_at_ball'[of r z0 UNIV] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1753
  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
  1754
    by eventually_elim (insert r, auto simp: field_simps)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1755
  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
  1756
    using r by (intro holomorphic_on_imp_continuous_on) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1757
  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
  1758
    by (auto simp: continuous_on_eq_continuous_at)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1759
  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
  1760
    unfolding isCont_def .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1761
  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
  1762
    by (blast intro: Lim_transform_eventually)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1763
  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
  1764
    by (rule filterlim_compose[OF _ g])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1765
  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
  1766
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1767
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1768
lemma zor_poly_pole_eqI:
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1769
  fixes f :: "complex \<Rightarrow> complex" and z0 :: complex
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1770
  defines "n \<equiv> zorder f z0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1771
  assumes f_iso:"isolated_singularity_at f z0" and "is_pole f z0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1772
  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
  1773
  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
  1774
  shows   "zor_poly f z0 z0 = c"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1775
proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1776
  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
  1777
  proof -
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1778
    have "\<exists>\<^sub>F w in at z0. f w \<noteq> 0"
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1779
      using non_zero_neighbour_pole[OF \<open>is_pole f z0\<close>] by (auto elim:eventually_frequentlyE)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1780
    moreover have "not_essential f z0" unfolding not_essential_def using \<open>is_pole f z0\<close> by simp
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1781
    ultimately show ?thesis using that zorder_exist[OF f_iso,folded n_def] by auto
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
  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
  1784
    using eventually_at_ball'[of r z0 UNIV] by auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1785
  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
  1786
    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
  1787
  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
  1788
    using r by (intro holomorphic_on_imp_continuous_on) auto
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1789
  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
  1790
    by (auto simp: continuous_on_eq_continuous_at)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1791
  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
  1792
    unfolding isCont_def .
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1793
  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
  1794
    by (blast intro: Lim_transform_eventually)
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1795
  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
  1796
    by (rule filterlim_compose[OF _ g])
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1797
  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
  1798
qed
6617fb368a06 Reorganised HOL-Complex_Analysis
Manuel Eberl <eberlm@in.tum.de>
parents:
diff changeset
  1799
77223
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1800
lemma
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1801
  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
  1802
  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
  1803
  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
  1804
    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
  1805
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1806
  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
  1807
    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
  1808
  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
  1809
    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
  1810
  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
  1811
    by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1812
  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
  1813
    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
  1814
             intro!: exI[of _ r] holomorphic_on_subset[OF holo])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1815
  hence ev: "\<forall>\<^sub>F w in at x. zor_poly f x w = f w * (w - x) ^ nat (- zorder f x)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1816
    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
  1817
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1818
  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
  1819
  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
  1820
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1821
  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"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1822
    "\<forall>w\<in>cball x r - {x}. f w = P w / (w - x) ^ n \<and> P w \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1823
    unfolding P_def n_def using zorder_exist_pole[OF holo assms(2,3,1)] by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1824
  have n: "n > 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1825
    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
  1826
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1827
  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
  1828
    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
  1829
    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
  1830
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1831
  define D where "D = (\<lambda>w. (deriv P w * (w - x) - of_nat n * P w) / (w - x) ^ (n + 1))"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1832
  define n' where "n' = n - 1"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1833
  have n': "n = Suc n'"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1834
    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
  1835
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1836
  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
  1837
    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
  1838
  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
  1839
    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
  1840
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1841
  {
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1842
    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
  1843
    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
  1844
      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
  1845
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1846
    have "((\<lambda>w. P w / (w - x) ^ n) 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
  1847
      apply (rule derivative_eq_intros refl | use w in force)+
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1848
      using w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1849
      apply (simp add: divide_simps D_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1850
      apply (simp add: n' algebra_simps)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1851
      done
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1852
    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
  1853
      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
  1854
    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
  1855
    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
  1856
      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
  1857
    ultimately have "D w = f' w"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1858
      using DERIV_unique by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1859
  } note D_eq = this
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1860
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1861
  have "is_pole D x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1862
    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
  1863
    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
  1864
  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
  1865
    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
  1866
  finally show "is_pole f' x" .
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1867
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1868
  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
  1869
  proof (rule zorder_eqI)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1870
    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
  1871
      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
  1872
    show "f' w = (deriv P w * (w - x) - of_nat n * P w) * (w - x) powr of_int (- int (Suc n))"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1873
      if "w \<in> ball x r" "w \<noteq> x" for w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1874
      using that D_eq[of w] n by (auto simp: D_def powr_diff powr_minus powr_nat' divide_simps)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1875
  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
  1876
  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
  1877
    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
  1878
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1879
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1880
lemma
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1881
  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
  1882
  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
  1883
    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
  1884
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1885
  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
  1886
    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
  1887
  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
  1888
    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
  1889
  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
  1890
    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
  1891
  show "is_pole (deriv f) x" "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
  1892
    by (rule is_pole_deriv' zorder_deriv', (rule assms * holomorphic_derivI holo | assumption)+)+
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1893
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1894
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1895
lemma removable_singularity_deriv':
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1896
  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
  1897
  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
  1898
  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
  1899
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1900
  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
  1901
    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
  1902
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1903
  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
  1904
  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
  1905
  proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1906
    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
  1907
      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
  1908
    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
  1909
      using assms that by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1910
    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
  1911
      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
  1912
    finally show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1913
      by (intro DERIV_imp_deriv assms)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1914
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1915
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1916
  have "g holomorphic_on A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1917
    unfolding g_def using assms assms(1) holo by (intro removable_singularity) auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1918
  hence "deriv g holomorphic_on A"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1919
    by (intro holomorphic_deriv assms)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1920
  hence "continuous_on A (deriv g)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1921
    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
  1922
  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
  1923
    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
  1924
  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
  1925
    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
  1926
  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
  1927
    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
  1928
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1929
    by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1930
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1931
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1932
lemma removable_singularity_deriv:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1933
  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
  1934
  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
  1935
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1936
  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
  1937
    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
  1938
  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
  1939
    using analytic_imp_holomorphic by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1940
  show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1941
    using assms(1)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1942
  proof (rule removable_singularity_deriv')
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1943
    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
  1944
      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
  1945
  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
  1946
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1947
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1948
lemma not_essential_deriv':
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1949
  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
  1950
  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
  1951
  shows   "not_essential f' x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1952
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1953
  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
  1954
    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
  1955
  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
  1956
    by (auto simp: not_essential_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1957
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1958
  proof cases
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1959
    case 1
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1960
    hence "is_pole f' x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1961
      using is_pole_deriv' assms by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1962
    thus ?thesis by (auto simp: not_essential_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1963
  next
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1964
    case (2 c)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1965
    from 2 have "\<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
  1966
      by (rule removable_singularity_deriv'[OF _ assms(2-4)])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1967
    thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1968
      by (auto simp: not_essential_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1969
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1970
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1971
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1972
lemma not_essential_deriv[singularity_intros]:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1973
  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
  1974
  shows   "not_essential (deriv f) x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1975
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1976
  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
  1977
    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
  1978
  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
  1979
    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
  1980
  show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1981
    using assms(1)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1982
  proof (rule not_essential_deriv')
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1983
    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
  1984
      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
  1985
  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
  1986
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1987
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1988
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
  1989
  fixes f :: "complex \<Rightarrow> complex"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1990
  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
  1991
  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
  1992
  shows   "f \<midarrow>z\<rightarrow> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1993
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1994
  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
  1995
    using frequently_atE by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1996
  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
  1997
    using g by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  1998
  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
  1999
    by (simp add: tendsto_eventually)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2000
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2001
  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
  2002
    unfolding not_essential_def by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2003
  thus ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2004
  proof cases
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2005
    case (1 c)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2006
    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
  2007
      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
  2008
    with fg have "c = 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2009
      using LIMSEQ_unique by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2010
    with 1 show ?thesis by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2011
  next
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2012
    case 2
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2013
    have "filterlim (\<lambda>x. f (g x)) at_infinity sequentially"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2014
      by (rule filterlim_compose[OF _ g(1)]) (use 2 in \<open>auto simp: is_pole_def\<close>)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2015
    with fg have False
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2016
      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
  2017
    thus ?thesis ..
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2018
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2019
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2020
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2021
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
  2022
  fixes f :: "complex \<Rightarrow> complex"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2023
  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
  2024
  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
  2025
  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
  2026
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2027
  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
  2028
    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
  2029
  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
  2030
    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
  2031
  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
  2032
    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
  2033
  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
  2034
    using frequently_eventually_frequently by blast
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2035
  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
  2036
    by (simp add: conj_commute)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2037
  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
  2038
    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
  2039
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2040
  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
  2041
  have "f \<midarrow>z\<rightarrow> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2042
    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
  2043
  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
  2044
    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
  2045
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2046
  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
  2047
  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
  2048
    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
  2049
      using r by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2050
    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
  2051
      by fact
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2052
    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
  2053
      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
  2054
  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
  2055
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2056
  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
  2057
    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
  2058
  thus "eventually (\<lambda>w. f w = 0) (at z)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2059
  proof eventually_elim
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2060
    case (elim w)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2061
    thus ?case using g_eq_0[of w]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2062
      by (auto simp: g_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2063
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2064
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2065
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2066
lemma pole_imp_not_constant:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2067
  fixes f :: "'a :: {perfect_space} \<Rightarrow> _"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2068
  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
  2069
  shows   "\<not>f constant_on B"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2070
proof
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2071
  assume *: "f constant_on B"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2072
  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
  2073
    by (auto simp: constant_on_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2074
  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
  2075
    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
  2076
  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
  2077
    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
  2078
  hence **: "f \<midarrow>x\<rightarrow> c"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2079
    by (simp add: tendsto_eventually)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2080
  show False
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2081
    using not_tendsto_and_filterlim_at_infinity[OF _ ** assms(1)[unfolded is_pole_def]] by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2082
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2083
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2084
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2085
lemma neg_zorder_imp_is_pole:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2086
  assumes iso:"isolated_singularity_at f z" and f_ness:"not_essential f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2087
      and "zorder f z < 0" and fre_nz:"\<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
  2088
    shows "is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2089
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2090
  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
  2091
  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
  2092
  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
  2093
  define nn where "nn = nat (-n)"
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
  obtain r where "P z \<noteq> 0" "r>0" and r_holo:"P holomorphic_on cball z r" and
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2096
       w_Pn:"(\<forall>w\<in>cball z r - {z}. f w = P w * (w - z) powr of_int n \<and> P w \<noteq> 0)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2097
    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
  2098
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2099
  have "is_pole (\<lambda>w. P w * (w - z) powr of_int n) z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2100
    unfolding is_pole_def
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2101
  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
  2102
    show "P \<midarrow>z\<rightarrow> P z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2103
      by (meson open_ball \<open>0 < r\<close> ball_subset_cball centre_in_ball
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2104
          continuous_on_eq_continuous_at continuous_on_subset
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2105
          holomorphic_on_imp_continuous_on isContD r_holo)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2106
    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
  2107
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2108
    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
  2109
      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
  2110
      using \<open>n<0\<close>
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2111
      by (auto intro!:tendsto_eq_intros filterlim_atI
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2112
              simp add:eventually_at_filter)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2113
    then show "LIM x at z. (x - z) powr of_int n :> at_infinity"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2114
    proof (elim filterlim_mono_eventually)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2115
      have "inverse ((x - z) ^ nat (-n)) = (x - z) powr of_int n"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2116
        if "x\<noteq>z" for x
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2117
        apply (subst powr_of_int)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2118
        using \<open>n<0\<close> using that by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2119
      then show "\<forall>\<^sub>F x in at z. inverse ((x - z) ^ nat (-n))
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2120
                  = (x - z) powr of_int n"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2121
        by (simp add: eventually_at_filter)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2122
    qed auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2123
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2124
  moreover have "\<forall>\<^sub>F w in at z. f w =  P w * (w - z) powr of_int n"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2125
    unfolding eventually_at_le
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2126
    apply (rule exI[where x=r])
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2127
    using w_Pn \<open>r>0\<close> by (simp add: dist_commute)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2128
  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
  2129
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2130
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2131
lemma is_pole_divide_zorder:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2132
  fixes f g::"complex \<Rightarrow> complex" and z::complex
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2133
  assumes f_iso:"isolated_singularity_at f z" and g_iso:"isolated_singularity_at g z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2134
      and f_ness:"not_essential f z" and g_ness:"not_essential g z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2135
      and fg_nconst: "\<exists>\<^sub>Fw in (at z). f w * g w\<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2136
      and z_less:"zorder f z < zorder g z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2137
    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
  2138
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2139
  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
  2140
                        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
  2141
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2142
  have "isolated_singularity_at fg z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2143
    unfolding fg_def using f_iso g_iso g_ness
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2144
    by (auto intro:singularity_intros)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2145
  moreover have "not_essential fg z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2146
    unfolding fg_def using f_iso g_iso g_ness f_ness
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2147
    by (auto intro:singularity_intros)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2148
  moreover have "zorder fg z < 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2149
  proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2150
    have "zorder fg z = fn - gn"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2151
      using zorder_divide[OF f_iso g_iso f_ness g_ness
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2152
            fg_nconst,folded fn_def gn_def fg_def] .
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2153
    then show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2154
      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
  2155
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2156
  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
  2157
    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
  2158
  ultimately show "is_pole fg z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2159
    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
  2160
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2161
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2162
lemma isolated_pole_imp_nzero_times:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2163
  assumes f_iso:"isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2164
    and "is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2165
  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
  2166
proof (rule ccontr)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2167
  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
  2168
  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
  2169
    unfolding not_frequently by simp
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2170
  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
  2171
    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
  2172
  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
  2173
    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
  2174
    .
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2175
  ultimately have "\<forall>\<^sub>F w in at z. False"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2176
    apply eventually_elim
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2177
    by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2178
  then show False by auto
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 isolated_pole_imp_neg_zorder:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2182
  assumes f_iso:"isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2183
    and "is_pole f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2184
  shows "zorder f z<0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2185
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2186
  obtain e where [simp]:"e>0" and f_holo:"f holomorphic_on ball z e - {z}"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2187
    using f_iso analytic_imp_holomorphic 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
  2188
  show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2189
    using zorder_exist_pole[OF f_holo,simplified,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
  2190
    by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2191
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2192
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2193
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
  2194
  assumes "isolated_singularity_at f x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2195
  shows "isolated_singularity_at (deriv f) x"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2196
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2197
  obtain r where "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
  2198
    using assms unfolding isolated_singularity_at_def by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2199
  from analytic_deriv[OF this(2)]
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2200
  have "deriv 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
  2201
  with \<open>r>0\<close> show ?thesis
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2202
    unfolding isolated_singularity_at_def by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2203
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2204
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2205
lemma zorder_deriv_minus_1:
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2206
  fixes f g::"complex \<Rightarrow> complex" and z::complex
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2207
  assumes f_iso:"isolated_singularity_at f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2208
      and f_ness:"not_essential f z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2209
      and f_nconst:"\<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
  2210
      and f_ord:"zorder f z \<noteq>0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2211
    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
  2212
proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2213
  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
  2214
  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
  2215
  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
  2216
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2217
  obtain r where "P z \<noteq> 0" "r>0" and P_holo:"P holomorphic_on cball z r"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2218
          and "(\<forall>w\<in>cball z r - {z}. f w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2219
                            = P w * (w - z) powr of_int n \<and> P w \<noteq> 0)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2220
    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
  2221
  from this(4)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2222
  have f_eq:"(\<forall>w\<in>cball z r - {z}. f w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2223
                            = P w * (w - z) powi n \<and> P w \<noteq> 0)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2224
    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
  2225
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2226
  define D where "D = (\<lambda>w. (deriv P w * (w - z) + of_int n * P w)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2227
                          * (w - z) powi (n - 1))"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2228
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2229
  have deriv_f_eq:"deriv f w = D w" 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
  2230
  proof -
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2231
    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
  2232
      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
  2233
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2234
    define wz where "wz = w - z"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2235
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2236
    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
  2237
    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
  2238
      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
  2239
          ball_subset_cball holomorphic_derivI holomorphic_on_subset that)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2240
    ultimately have "((\<lambda>w. P w * (w - z) powi n) 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
  2241
      unfolding D_def using that
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2242
      apply (auto intro!: derivative_eq_intros)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2243
      apply (fold wz_def)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2244
      by (auto simp:algebra_simps simp flip:power_int_add_1')
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2245
    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
  2246
      using f_eq
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2247
      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
  2248
    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
  2249
    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
  2250
      by (metis DERIV_imp_deriv calculation)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2251
    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
  2252
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2253
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2254
  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
  2255
  proof (rule zorder_eqI)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2256
    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
  2257
      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
  2258
    define g where "g=(\<lambda>w. (deriv P w * (w - z) + of_int n * P w))"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2259
    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
  2260
      unfolding g_def using P_holo
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2261
      by (auto intro!:holomorphic_intros)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2262
    show "g z \<noteq> 0"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2263
      unfolding g_def using \<open>P z \<noteq> 0\<close> \<open>n\<noteq>0\<close> by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2264
    show "deriv f w =
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2265
         (deriv P w * (w - z) + of_int n * P w) * (w - z) powr of_int (n - 1)"
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2266
      if "w \<in> ball z r" "w \<noteq> z" for w
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2267
      apply (subst complex_powr_of_int)
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2268
      using deriv_f_eq that unfolding D_def by auto
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2269
  qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2270
qed
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2271
607e1e345e8f Lots of new material chiefly about complex analysis
paulson <lp15@cam.ac.uk>
parents: 76900
diff changeset
  2272
end