| author | haftmann | 
| Sun, 26 Nov 2023 21:06:46 +0000 | |
| changeset 79070 | a4775fe69f5d | 
| parent 78698 | 1b9388e6eb75 | 
| child 79857 | 819c28a7280f | 
| permissions | -rw-r--r-- | 
| 77277 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1 | theory Meromorphic | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2 | imports Laurent_Convergence Riemann_Mapping | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 3 | begin | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 4 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 5 | definition remove_sings :: "(complex \<Rightarrow> complex) \<Rightarrow> complex \<Rightarrow> complex" where | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 6 | "remove_sings f z = (if \<exists>c. f \<midarrow>z\<rightarrow> c then Lim (at z) f else 0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 7 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 8 | lemma remove_sings_eqI [intro]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 9 | assumes "f \<midarrow>z\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 10 | shows "remove_sings f z = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 11 | using assms unfolding remove_sings_def by (auto simp: tendsto_Lim) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 12 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 13 | lemma remove_sings_at_analytic [simp]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 14 |   assumes "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 15 | shows "remove_sings f z = f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 16 | using assms by (intro remove_sings_eqI) (simp add: analytic_at_imp_isCont isContD) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 17 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 18 | lemma remove_sings_at_pole [simp]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 19 | assumes "is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 20 | shows "remove_sings f z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 21 | using assms unfolding remove_sings_def is_pole_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 22 | by (meson at_neq_bot not_tendsto_and_filterlim_at_infinity) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 23 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 24 | lemma eventually_remove_sings_eq_at: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 25 | assumes "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 26 | shows "eventually (\<lambda>w. remove_sings f w = f w) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 27 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 28 |   from assms obtain r where r: "r > 0" "f analytic_on ball z r - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 29 | by (auto simp: isolated_singularity_at_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 30 |   hence *: "f analytic_on {w}" if "w \<in> ball z r - {z}" for w
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 31 | using r that by (auto intro: analytic_on_subset) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 32 |   have "eventually (\<lambda>w. w \<in> ball z r - {z}) (at z)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 33 | using r by (intro eventually_at_in_open) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 34 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 35 | by eventually_elim (auto simp: remove_sings_at_analytic *) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 36 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 37 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 38 | lemma eventually_remove_sings_eq_nhds: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 39 |   assumes "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 40 | shows "eventually (\<lambda>w. remove_sings f w = f w) (nhds z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 41 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 42 | from assms obtain A where A: "open A" "z \<in> A" "f holomorphic_on A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 43 | by (auto simp: analytic_at) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 44 | have "eventually (\<lambda>z. z \<in> A) (nhds z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 45 | by (intro eventually_nhds_in_open A) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 46 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 47 | proof eventually_elim | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 48 | case (elim w) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 49 |     from elim have "f analytic_on {w}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 50 | using A analytic_at by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 51 | thus ?case by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 52 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 53 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 54 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 55 | lemma remove_sings_compose: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 56 | assumes "filtermap g (at z) = at z'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 57 | shows "remove_sings (f \<circ> g) z = remove_sings f z'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 58 | proof (cases "\<exists>c. f \<midarrow>z'\<rightarrow> c") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 59 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 60 | then obtain c where c: "f \<midarrow>z'\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 61 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 62 | from c have "remove_sings f z' = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 63 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 64 | moreover from c have "remove_sings (f \<circ> g) z = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 65 | using c by (intro remove_sings_eqI) (auto simp: filterlim_def filtermap_compose assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 66 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 67 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 68 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 69 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 70 | hence "\<not>(\<exists>c. (f \<circ> g) \<midarrow>z\<rightarrow> c)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 71 | by (auto simp: filterlim_def filtermap_compose assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 72 | with False show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 73 | by (auto simp: remove_sings_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 74 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 75 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 76 | lemma remove_sings_cong: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 77 | assumes "eventually (\<lambda>x. f x = g x) (at z)" "z = z'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 78 | shows "remove_sings f z = remove_sings g z'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 79 | proof (cases "\<exists>c. f \<midarrow>z\<rightarrow> c") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 80 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 81 | then obtain c where c: "f \<midarrow>z\<rightarrow> c" by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 82 | hence "remove_sings f z = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 83 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 84 | moreover have "f \<midarrow>z\<rightarrow> c \<longleftrightarrow> g \<midarrow>z'\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 85 | using assms by (intro filterlim_cong refl) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 86 | with c have "remove_sings g z' = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 87 | by (intro remove_sings_eqI) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 88 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 89 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 90 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 91 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 92 | have "f \<midarrow>z\<rightarrow> c \<longleftrightarrow> g \<midarrow>z'\<rightarrow> c" for c | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 93 | using assms by (intro filterlim_cong) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 94 | with False show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 95 | by (auto simp: remove_sings_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 96 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 97 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 98 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 99 | lemma deriv_remove_sings_at_analytic [simp]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 100 |   assumes "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 101 | shows "deriv (remove_sings f) z = deriv f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 102 | apply (rule deriv_cong_ev) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 103 | apply (rule eventually_remove_sings_eq_nhds) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 104 | using assms by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 105 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 106 | lemma isolated_singularity_at_remove_sings [simp, intro]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 107 | assumes "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 108 | shows "isolated_singularity_at (remove_sings f) z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 109 | using isolated_singularity_at_cong[OF eventually_remove_sings_eq_at[OF assms] refl] assms | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 110 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 111 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 112 | lemma not_essential_remove_sings_iff [simp]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 113 | assumes "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 114 | shows "not_essential (remove_sings f) z \<longleftrightarrow> not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 115 | using not_essential_cong[OF eventually_remove_sings_eq_at[OF assms(1)] refl] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 116 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 117 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 118 | lemma not_essential_remove_sings [intro]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 119 | assumes "isolated_singularity_at f z" "not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 120 | shows "not_essential (remove_sings f) z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 121 | by (subst not_essential_remove_sings_iff) (use assms in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 122 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 123 | lemma | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 124 | assumes "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 125 | shows is_pole_remove_sings_iff [simp]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 126 | "is_pole (remove_sings f) z \<longleftrightarrow> is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 127 | and zorder_remove_sings [simp]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 128 | "zorder (remove_sings f) z = zorder f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 129 | and zor_poly_remove_sings [simp]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 130 | "zor_poly (remove_sings f) z = zor_poly f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 131 | and has_laurent_expansion_remove_sings_iff [simp]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 132 | "(\<lambda>w. remove_sings f (z + w)) has_laurent_expansion F \<longleftrightarrow> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 133 | (\<lambda>w. f (z + w)) has_laurent_expansion F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 134 | and tendsto_remove_sings_iff [simp]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 135 | "remove_sings f \<midarrow>z\<rightarrow> c \<longleftrightarrow> f \<midarrow>z\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 136 | by (intro is_pole_cong eventually_remove_sings_eq_at refl zorder_cong | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 137 | zor_poly_cong has_laurent_expansion_cong' tendsto_cong assms)+ | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 138 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 139 | lemma get_all_poles_from_remove_sings: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 140 | fixes f:: "complex \<Rightarrow> complex" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 141 | defines "ff\<equiv>remove_sings f" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 142 | assumes f_holo:"f holomorphic_on s - pts" and "finite pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 143 | "pts\<subseteq>s" "open s" and not_ess:"\<forall>x\<in>pts. not_essential f x" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 144 | obtains pts' where | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 145 | "pts' \<subseteq> pts" "finite pts'" "ff holomorphic_on s - pts'" "\<forall>x\<in>pts'. is_pole ff x" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 146 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 147 |   define pts' where "pts' = {x\<in>pts. is_pole f x}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 148 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 149 | have "pts' \<subseteq> pts" unfolding pts'_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 150 | then have "finite pts'" using \<open>finite pts\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 151 | using rev_finite_subset by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 152 | then have "open (s - pts')" using \<open>open s\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 153 | by (simp add: finite_imp_closed open_Diff) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 154 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 155 | have isolated:"isolated_singularity_at f z" if "z\<in>pts" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 156 | proof (rule isolated_singularity_at_holomorphic) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 157 |     show "f holomorphic_on (s-(pts-{z})) - {z}" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 158 | by (metis Diff_insert f_holo insert_Diff that) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 159 |     show " open (s - (pts - {z}))" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 160 | by (meson assms(3) assms(5) finite_Diff finite_imp_closed open_Diff) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 161 |     show "z \<in> s - (pts - {z})" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 162 | using assms(4) that by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 163 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 164 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 165 | have "ff holomorphic_on s - pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 166 | proof (rule no_isolated_singularity') | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 167 | show "(ff \<longlongrightarrow> ff z) (at z within s - pts')" if "z \<in> pts-pts'" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 168 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 169 | have "at z within s - pts' = at z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 170 | apply (rule at_within_open) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 171 | using \<open>open (s - pts')\<close> that \<open>pts\<subseteq>s\<close> by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 172 | moreover have "ff \<midarrow>z\<rightarrow> ff z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 173 | unfolding ff_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 174 | proof (subst tendsto_remove_sings_iff) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 175 | show "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 176 | apply (rule isolated) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 177 | using that by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 178 | have "not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 179 | using not_ess that by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 180 | moreover have "\<not>is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 181 | using that unfolding pts'_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 182 | ultimately have "\<exists>c. f \<midarrow>z\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 183 | unfolding not_essential_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 184 | then show "f \<midarrow>z\<rightarrow> remove_sings f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 185 | using remove_sings_eqI by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 186 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 187 | ultimately show ?thesis by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 188 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 189 | have "ff holomorphic_on s - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 190 | using f_holo | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 191 | proof (elim holomorphic_transform) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 192 | fix x assume "x \<in> s - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 193 |       then have "f analytic_on {x}" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 194 | using assms(3) assms(5) f_holo | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 195 | by (meson finite_imp_closed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 196 | holomorphic_on_imp_analytic_at open_Diff) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 197 | from remove_sings_at_analytic[OF this] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 198 | show "f x = ff x" unfolding ff_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 199 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 200 | then show "ff holomorphic_on s - pts' - (pts - pts')" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 201 | apply (elim holomorphic_on_subset) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 202 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 203 | show "open (s - pts')" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 204 | by (simp add: \<open>open (s - pts')\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 205 | show "finite (pts - pts')" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 206 | by (simp add: assms(3)) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 207 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 208 | moreover have "\<forall>x\<in>pts'. is_pole ff x" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 209 | unfolding pts'_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 210 | using ff_def is_pole_remove_sings_iff isolated by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 211 | moreover note \<open>pts' \<subseteq> pts\<close> \<open>finite pts'\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 212 | ultimately show ?thesis using that by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 213 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 214 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 215 | lemma remove_sings_eq_0_iff: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 216 | assumes "not_essential f w" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 217 | shows "remove_sings f w = 0 \<longleftrightarrow> is_pole f w \<or> f \<midarrow>w\<rightarrow> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 218 | proof (cases "is_pole f w") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 219 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 220 | then show ?thesis by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 221 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 222 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 223 | then obtain c where c:"f \<midarrow>w\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 224 | using \<open>not_essential f w\<close> unfolding not_essential_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 225 | then show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 226 | using False remove_sings_eqI by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 227 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 228 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 229 | definition meromorphic_on:: "[complex \<Rightarrow> complex, complex set, complex set] \<Rightarrow> bool" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 230 |   ("_ (meromorphic'_on) _ _" [50,50,50]50) where 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 231 | "f meromorphic_on D pts \<equiv> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 232 | open D \<and> pts \<subseteq> D \<and> (\<forall>z\<in>pts. isolated_singularity_at f z \<and> not_essential f z) \<and> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 233 | (\<forall>z\<in>D. \<not>(z islimpt pts)) \<and> (f holomorphic_on D-pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 234 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 235 | lemma meromorphic_imp_holomorphic: "f meromorphic_on D pts \<Longrightarrow> f holomorphic_on (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 236 | unfolding meromorphic_on_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 237 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 238 | lemma meromorphic_imp_closedin_pts: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 239 | assumes "f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 240 | shows "closedin (top_of_set D) pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 241 | by (meson assms closedin_limpt meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 242 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 243 | lemma meromorphic_imp_open_diff': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 244 | assumes "f meromorphic_on D pts" "pts' \<subseteq> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 245 | shows "open (D - pts')" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 246 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 247 | have "D - pts' = D - closure pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 248 | proof safe | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 249 | fix x assume x: "x \<in> D" "x \<in> closure pts'" "x \<notin> pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 250 | hence "x islimpt pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 251 | by (subst islimpt_in_closure) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 252 | hence "x islimpt pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 253 | by (rule islimpt_subset) fact | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 254 | with assms x show False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 255 | by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 256 | qed (use closure_subset in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 257 | then show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 258 | using assms meromorphic_on_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 259 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 260 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 261 | lemma meromorphic_imp_open_diff: "f meromorphic_on D pts \<Longrightarrow> open (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 262 | by (erule meromorphic_imp_open_diff') auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 263 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 264 | lemma meromorphic_pole_subset: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 265 | assumes merf: "f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 266 |   shows "{x\<in>D. is_pole f x} \<subseteq> pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 267 | by (smt (verit) Diff_iff assms mem_Collect_eq meromorphic_imp_open_diff | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 268 | meromorphic_on_def not_is_pole_holomorphic subsetI) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 269 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 270 | named_theorems meromorphic_intros | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 271 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 272 | lemma meromorphic_on_subset: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 273 | assumes "f meromorphic_on A pts" "open B" "B \<subseteq> A" "pts' = pts \<inter> B" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 274 | shows "f meromorphic_on B pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 275 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 276 | proof (intro ballI conjI) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 277 | fix z assume "z \<in> B" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 278 | show "\<not>z islimpt pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 279 | proof | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 280 | assume "z islimpt pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 281 | hence "z islimpt pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 282 | by (rule islimpt_subset) (use \<open>pts' = _\<close> in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 283 | thus False using \<open>z \<in> B\<close> \<open>B \<subseteq> A\<close> assms(1) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 284 | by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 285 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 286 | qed (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 287 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 288 | lemma meromorphic_on_superset_pts: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 289 | assumes "f meromorphic_on A pts" "pts \<subseteq> pts'" "pts' \<subseteq> A" "\<forall>x\<in>A. \<not>x islimpt pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 290 | shows "f meromorphic_on A pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 291 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 292 | proof (intro conjI ballI impI) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 293 | fix z assume "z \<in> pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 294 | from assms(1) have holo: "f holomorphic_on A - pts" and "open A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 295 | unfolding meromorphic_on_def by blast+ | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 296 | have "open (A - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 297 | by (intro meromorphic_imp_open_diff[OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 298 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 299 | show "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 300 | proof (cases "z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 301 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 302 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 303 | using \<open>open (A - pts)\<close> assms \<open>z \<in> pts'\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 304 | by (intro isolated_singularity_at_holomorphic[of _ "A - pts"] holomorphic_on_subset[OF holo]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 305 | auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 306 | qed (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 307 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 308 | show "not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 309 | proof (cases "z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 310 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 311 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 312 | using \<open>open (A - pts)\<close> assms \<open>z \<in> pts'\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 313 | by (intro not_essential_holomorphic[of _ "A - pts"] holomorphic_on_subset[OF holo]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 314 | auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 315 | qed (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 316 | qed (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 317 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 318 | lemma meromorphic_on_no_singularities: "f meromorphic_on A {} \<longleftrightarrow> f holomorphic_on A \<and> open A"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 319 | by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 320 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 321 | lemma holomorphic_on_imp_meromorphic_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 322 | "f holomorphic_on A \<Longrightarrow> pts \<subseteq> A \<Longrightarrow> open A \<Longrightarrow> \<forall>x\<in>A. \<not>x islimpt pts \<Longrightarrow> f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 323 |   by (rule meromorphic_on_superset_pts[where pts = "{}"])
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 324 | (auto simp: meromorphic_on_no_singularities) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 325 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 326 | lemma meromorphic_on_const [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 327 | assumes "open A" "\<forall>x\<in>A. \<not>x islimpt pts" "pts \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 328 | shows "(\<lambda>_. c) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 329 | by (rule holomorphic_on_imp_meromorphic_on) (use assms in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 330 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 331 | lemma meromorphic_on_ident [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 332 | assumes "open A" "\<forall>x\<in>A. \<not>x islimpt pts" "pts \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 333 | shows "(\<lambda>x. x) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 334 | by (rule holomorphic_on_imp_meromorphic_on) (use assms in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 335 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 336 | lemma meromorphic_on_id [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 337 | assumes "open A" "\<forall>x\<in>A. \<not>x islimpt pts" "pts \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 338 | shows "id meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 339 | using meromorphic_on_ident assms unfolding id_def . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 340 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 341 | lemma not_essential_add [singularity_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 342 | assumes f_ness: "not_essential f z" and g_ness: "not_essential g z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 343 | assumes f_iso: "isolated_singularity_at f z" and g_iso: "isolated_singularity_at g z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 344 | shows "not_essential (\<lambda>w. f w + g w) z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 345 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 346 | have "(\<lambda>w. f (z + w) + g (z + w)) has_laurent_expansion laurent_expansion f z + laurent_expansion g z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 347 | by (intro not_essential_has_laurent_expansion laurent_expansion_intros assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 348 | hence "not_essential (\<lambda>w. f (z + w) + g (z + w)) 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 349 | using has_laurent_expansion_not_essential_0 by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 350 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 351 | by (simp add: not_essential_shift_0) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 352 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 353 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 354 | lemma meromorphic_on_uminus [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 355 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 356 | shows "(\<lambda>z. -f z) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 357 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 358 | by (use assms in \<open>auto simp: meromorphic_on_def intro!: holomorphic_intros singularity_intros\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 359 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 360 | lemma meromorphic_on_add [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 361 | assumes "f meromorphic_on A pts" "g meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 362 | shows "(\<lambda>z. f z + g z) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 363 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 364 | by (use assms in \<open>auto simp: meromorphic_on_def intro!: holomorphic_intros singularity_intros\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 365 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 366 | lemma meromorphic_on_add': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 367 | assumes "f meromorphic_on A pts1" "g meromorphic_on A pts2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 368 | shows "(\<lambda>z. f z + g z) meromorphic_on A (pts1 \<union> pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 369 | proof (rule meromorphic_intros) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 370 | show "f meromorphic_on A (pts1 \<union> pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 371 | by (rule meromorphic_on_superset_pts[OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 372 | (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 373 | show "g meromorphic_on A (pts1 \<union> pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 374 | by (rule meromorphic_on_superset_pts[OF assms(2)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 375 | (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 376 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 377 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 378 | lemma meromorphic_on_add_const [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 379 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 380 | shows "(\<lambda>z. f z + c) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 381 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 382 | by (use assms in \<open>auto simp: meromorphic_on_def intro!: holomorphic_intros singularity_intros\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 383 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 384 | lemma meromorphic_on_minus_const [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 385 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 386 | shows "(\<lambda>z. f z - c) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 387 | using meromorphic_on_add_const[OF assms,of "-c"] by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 388 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 389 | lemma meromorphic_on_diff [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 390 | assumes "f meromorphic_on A pts" "g meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 391 | shows "(\<lambda>z. f z - g z) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 392 | using meromorphic_on_add[OF assms(1) meromorphic_on_uminus[OF assms(2)]] by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 393 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 394 | lemma meromorphic_on_diff': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 395 | assumes "f meromorphic_on A pts1" "g meromorphic_on A pts2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 396 | shows "(\<lambda>z. f z - g z) meromorphic_on A (pts1 \<union> pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 397 | proof (rule meromorphic_intros) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 398 | show "f meromorphic_on A (pts1 \<union> pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 399 | by (rule meromorphic_on_superset_pts[OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 400 | (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 401 | show "g meromorphic_on A (pts1 \<union> pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 402 | by (rule meromorphic_on_superset_pts[OF assms(2)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 403 | (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 404 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 405 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 406 | lemma meromorphic_on_mult [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 407 | assumes "f meromorphic_on A pts" "g meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 408 | shows "(\<lambda>z. f z * g z) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 409 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 410 | by (use assms in \<open>auto simp: meromorphic_on_def intro!: holomorphic_intros singularity_intros\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 411 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 412 | lemma meromorphic_on_mult': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 413 | assumes "f meromorphic_on A pts1" "g meromorphic_on A pts2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 414 | shows "(\<lambda>z. f z * g z) meromorphic_on A (pts1 \<union> pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 415 | proof (rule meromorphic_intros) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 416 | show "f meromorphic_on A (pts1 \<union> pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 417 | by (rule meromorphic_on_superset_pts[OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 418 | (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 419 | show "g meromorphic_on A (pts1 \<union> pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 420 | by (rule meromorphic_on_superset_pts[OF assms(2)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 421 | (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 422 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 423 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 424 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 425 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 426 | lemma meromorphic_on_imp_not_essential: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 427 | assumes "f meromorphic_on A pts" "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 428 | shows "not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 429 | proof (cases "z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 430 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 431 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 432 | using not_essential_holomorphic[of f "A - pts" z] meromorphic_imp_open_diff[OF assms(1)] assms | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 433 | by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 434 | qed (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 435 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 436 | lemma meromorphic_imp_analytic: "f meromorphic_on D pts \<Longrightarrow> f analytic_on (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 437 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 438 | apply (subst analytic_on_open) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 439 | using meromorphic_imp_open_diff meromorphic_on_id apply blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 440 | apply auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 441 | done | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 442 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 443 | lemma not_islimpt_isolated_zeros: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 444 | assumes mero: "f meromorphic_on A pts" and "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 445 |   shows "\<not>z islimpt {w\<in>A. isolated_zero f w}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 446 | proof | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 447 |   assume islimpt: "z islimpt {w\<in>A. isolated_zero f w}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 448 | have holo: "f holomorphic_on A - pts" and "open A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 449 | using assms by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 450 |   have open': "open (A - (pts - {z}))"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 451 | by (intro meromorphic_imp_open_diff'[OF mero]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 452 |   then obtain r where r: "r > 0" "ball z r \<subseteq> A - (pts - {z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 453 | using meromorphic_imp_open_diff[OF mero] \<open>z \<in> A\<close> openE by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 454 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 455 | have "not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 456 | using assms by (rule meromorphic_on_imp_not_essential) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 457 | then consider c where "f \<midarrow>z\<rightarrow> c" | "is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 458 | unfolding not_essential_def by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 459 | thus False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 460 | proof cases | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 461 | assume "is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 462 | hence "eventually (\<lambda>w. f w \<noteq> 0) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 463 | by (rule non_zero_neighbour_pole) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 464 |     hence "\<not>z islimpt {w. f w = 0}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 465 | by (simp add: islimpt_conv_frequently_at frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 466 |     moreover have "z islimpt {w. f w = 0}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 467 | using islimpt by (rule islimpt_subset) (auto simp: isolated_zero_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 468 | ultimately show False by contradiction | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 469 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 470 | fix c assume c: "f \<midarrow>z\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 471 | define g where "g = (\<lambda>w. if w = z then c else f w)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 472 |     have holo': "g holomorphic_on A - (pts - {z})" unfolding g_def
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 473 | by (intro removable_singularity holomorphic_on_subset[OF holo] open' c) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 474 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 475 | have eq_zero: "g w = 0" if "w \<in> ball z r" for w | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 476 | proof (rule analytic_continuation[where f = g]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 477 |       show "open (ball z r)" "connected (ball z r)" "{w\<in>ball z r. isolated_zero f w} \<subseteq> ball z r"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 478 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 479 |       have "z islimpt {w\<in>A. isolated_zero f w} \<inter> ball z r"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 480 | using islimpt \<open>r > 0\<close> by (intro islimpt_Int_eventually eventually_at_in_open') auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 481 |       also have "\<dots> = {w\<in>ball z r. isolated_zero f w}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 482 | using r by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 483 |       finally show "z islimpt {w\<in>ball z r. isolated_zero f w}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 484 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 485 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 486 |       fix w assume w: "w \<in> {w\<in>ball z r. isolated_zero f w}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 487 | show "g w = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 488 | proof (cases "w = z") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 489 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 490 | thus ?thesis using w by (auto simp: g_def isolated_zero_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 491 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 492 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 493 |         have "z islimpt {z. f z = 0}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 494 | using islimpt by (rule islimpt_subset) (auto simp: isolated_zero_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 495 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 496 | using w by (simp add: isolated_zero_altdef True) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 497 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 498 | qed (use r that in \<open>auto intro!: holomorphic_on_subset[OF holo'] simp: isolated_zero_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 499 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 500 |     have "infinite ({w\<in>A. isolated_zero f w} \<inter> ball z r)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 501 | using islimpt \<open>r > 0\<close> unfolding islimpt_eq_infinite_ball by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 502 |     hence "{w\<in>A. isolated_zero f w} \<inter> ball z r \<noteq> {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 503 | by force | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 504 | then obtain z0 where z0: "z0 \<in> A" "isolated_zero f z0" "z0 \<in> ball z r" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 505 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 506 |     have "\<forall>\<^sub>F y in at z0. y \<in> ball z r - (if z = z0 then {} else {z}) - {z0}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 507 | using r z0 by (intro eventually_at_in_open) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 508 | hence "eventually (\<lambda>w. f w = 0) (at z0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 509 | proof eventually_elim | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 510 | case (elim w) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 511 | show ?case | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 512 | using eq_zero[of w] elim by (auto simp: g_def split: if_splits) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 513 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 514 | hence "eventually (\<lambda>w. f w = 0) (at z0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 515 | by (auto simp: g_def eventually_at_filter elim!: eventually_mono split: if_splits) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 516 | moreover from z0 have "eventually (\<lambda>w. f w \<noteq> 0) (at z0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 517 | by (simp add: isolated_zero_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 518 | ultimately have "eventually (\<lambda>_. False) (at z0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 519 | by eventually_elim auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 520 | thus False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 521 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 522 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 523 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 524 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 525 | lemma closedin_isolated_zeros: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 526 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 527 |   shows   "closedin (top_of_set A) {z\<in>A. isolated_zero f z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 528 | unfolding closedin_limpt using not_islimpt_isolated_zeros[OF assms] by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 529 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 530 | lemma meromorphic_on_deriv': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 531 | assumes "f meromorphic_on A pts" "open A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 532 | assumes "\<And>x. x \<in> A - pts \<Longrightarrow> (f has_field_derivative f' x) (at x)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 533 | shows "f' meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 534 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 535 | proof (intro conjI ballI) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 536 | have "open (A - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 537 | by (intro meromorphic_imp_open_diff[OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 538 | thus "f' holomorphic_on A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 539 | by (rule derivative_is_holomorphic) (use assms in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 540 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 541 | fix z assume "z \<in> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 542 | hence "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 543 | using assms(1) by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 544 |   from \<open>z \<in> pts\<close> obtain r where r: "r > 0" "f analytic_on ball z r - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 545 | using assms(1) by (auto simp: meromorphic_on_def isolated_singularity_at_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 546 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 547 |   have "open (ball z r \<inter> (A - (pts - {z})))"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 548 | by (intro open_Int assms meromorphic_imp_open_diff'[OF assms(1)]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 549 |   then obtain r' where r': "r' > 0" "ball z r' \<subseteq> ball z r \<inter> (A - (pts - {z}))"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 550 | using r \<open>z \<in> A\<close> by (subst (asm) open_contains_ball) fastforce | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 551 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 552 |   have "open (ball z r' - {z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 553 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 554 |   hence "f' holomorphic_on ball z r' - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 555 | by (rule derivative_is_holomorphic[of _ f]) (use r' in \<open>auto intro!: assms(3)\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 556 |   moreover have "open (ball z r' - {z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 557 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 558 | ultimately show "isolated_singularity_at f' z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 559 | unfolding isolated_singularity_at_def using \<open>r' > 0\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 560 | by (auto simp: analytic_on_open intro!: exI[of _ r']) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 561 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 562 | fix z assume z: "z \<in> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 563 | hence z': "not_essential f z" "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 564 | using assms by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 565 | from z'(1) show "not_essential f' z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 566 | proof (rule not_essential_deriv') | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 567 |     show "z \<in> A - (pts - {z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 568 | using \<open>z \<in> A\<close> by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 569 |     show "open (A - (pts - {z}))"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 570 | by (intro meromorphic_imp_open_diff'[OF assms(1)]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 571 | qed (use assms in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 572 | qed (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 573 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 574 | lemma meromorphic_on_deriv [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 575 | assumes "f meromorphic_on A pts" "open A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 576 | shows "deriv f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 577 | proof (intro meromorphic_on_deriv'[OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 578 | have *: "open (A - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 579 | by (intro meromorphic_imp_open_diff[OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 580 | show "(f has_field_derivative deriv f x) (at x)" if "x \<in> A - pts" for x | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 581 | using assms(1) by (intro holomorphic_derivI[OF _ * that]) (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 582 | qed fact | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 583 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 584 | lemma meromorphic_on_imp_analytic_at: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 585 | assumes "f meromorphic_on A pts" "z \<in> A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 586 |   shows   "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 587 | using assms by (metis analytic_at meromorphic_imp_open_diff meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 588 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 589 | lemma meromorphic_compact_finite_pts: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 590 | assumes "f meromorphic_on D pts" "compact S" "S \<subseteq> D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 591 | shows "finite (S \<inter> pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 592 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 593 |   { assume "infinite (S \<inter> pts)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 594 | then obtain z where "z \<in> S" and z: "z islimpt (S \<inter> pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 595 | using assms by (metis compact_eq_Bolzano_Weierstrass inf_le1) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 596 | then have False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 597 | using assms by (meson in_mono inf_le2 islimpt_subset meromorphic_on_def) } | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 598 | then show ?thesis by metis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 599 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 600 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 601 | lemma meromorphic_imp_countable: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 602 | assumes "f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 603 | shows "countable pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 604 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 605 | obtain K :: "nat \<Rightarrow> complex set" where K: "D = (\<Union>n. K n)" "\<And>n. compact (K n)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 606 | using assms unfolding meromorphic_on_def by (metis open_Union_compact_subsets) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 607 | then have "pts = (\<Union>n. K n \<inter> pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 608 | using assms meromorphic_on_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 609 | moreover have "\<And>n. finite (K n \<inter> pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 610 | by (metis K(1) K(2) UN_I assms image_iff meromorphic_compact_finite_pts rangeI subset_eq) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 611 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 612 | by (metis countableI_type countable_UN countable_finite) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 613 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 614 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 615 | lemma meromorphic_imp_connected_diff': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 616 | assumes "f meromorphic_on D pts" "connected D" "pts' \<subseteq> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 617 | shows "connected (D - pts')" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 618 | proof (rule connected_open_diff_countable) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 619 | show "countable pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 620 | by (rule countable_subset [OF assms(3)]) (use assms(1) in \<open>auto simp: meromorphic_imp_countable\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 621 | qed (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 622 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 623 | lemma meromorphic_imp_connected_diff: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 624 | assumes "f meromorphic_on D pts" "connected D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 625 | shows "connected (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 626 | using meromorphic_imp_connected_diff'[OF assms order.refl] . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 627 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 628 | lemma meromorphic_on_compose [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 629 | assumes f: "f meromorphic_on A pts" and g: "g holomorphic_on B" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 630 | assumes "open B" and "g ` B \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 631 | shows "(\<lambda>x. f (g x)) meromorphic_on B (isolated_points_of (g -` pts \<inter> B))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 632 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 633 | proof (intro ballI conjI) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 634 | fix z assume z: "z \<in> isolated_points_of (g -` pts \<inter> B)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 635 | hence z': "z \<in> B" "g z \<in> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 636 | using isolated_points_of_subset by blast+ | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 637 |   have g': "g analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 638 | using g z' \<open>open B\<close> analytic_at by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 639 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 640 | show "isolated_singularity_at (\<lambda>x. f (g x)) z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 641 | by (rule isolated_singularity_at_compose[OF _ g']) (use f z' in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 642 | show "not_essential (\<lambda>x. f (g x)) z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 643 | by (rule not_essential_compose[OF _ g']) (use f z' in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 644 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 645 | fix z assume z: "z \<in> B" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 646 | hence "g z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 647 | using assms by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 648 | hence "\<not>g z islimpt pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 649 | using f by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 650 | hence ev: "eventually (\<lambda>w. w \<notin> pts) (at (g z))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 651 | by (auto simp: islimpt_conv_frequently_at frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 652 |   have g': "g analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 653 | by (rule holomorphic_on_imp_analytic_at[OF g]) (use assms z in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 654 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 655 | (* TODO: There's probably a useful lemma somewhere in here to extract... *) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 656 | have "eventually (\<lambda>w. w \<notin> isolated_points_of (g -` pts \<inter> B)) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 657 | proof (cases "isolated_zero (\<lambda>w. g w - g z) z") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 658 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 659 | have "eventually (\<lambda>w. w \<notin> pts) (at (g z))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 660 | using ev by (auto simp: islimpt_conv_frequently_at frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 661 | moreover have "g \<midarrow>z\<rightarrow> g z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 662 | using analytic_at_imp_isCont[OF g'] isContD by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 663 | hence lim: "filterlim g (at (g z)) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 664 | using True by (auto simp: filterlim_at isolated_zero_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 665 | have "eventually (\<lambda>w. g w \<notin> pts) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 666 | using ev lim by (rule eventually_compose_filterlim) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 667 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 668 | by eventually_elim (auto simp: isolated_points_of_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 669 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 670 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 671 | have "eventually (\<lambda>w. g w - g z = 0) (nhds z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 672 | using False by (rule non_isolated_zero) (auto intro!: analytic_intros g') | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 673 | hence "eventually (\<lambda>w. g w = g z \<and> w \<in> B) (nhds z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 674 | using eventually_nhds_in_open[OF \<open>open B\<close> \<open>z \<in> B\<close>] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 675 | by eventually_elim auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 676 | then obtain X where X: "open X" "z \<in> X" "X \<subseteq> B" "\<forall>x\<in>X. g x = g z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 677 | unfolding eventually_nhds by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 678 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 679 | have "z0 \<notin> isolated_points_of (g -` pts \<inter> B)" if "z0 \<in> X" for z0 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 680 | proof (cases "g z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 681 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 682 | with that have "g z0 \<notin> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 683 | using X by metis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 684 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 685 | by (auto simp: isolated_points_of_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 686 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 687 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 688 | have "eventually (\<lambda>w. w \<in> X) (at z0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 689 | by (intro eventually_at_in_open') fact+ | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 690 | hence "eventually (\<lambda>w. w \<in> g -` pts \<inter> B) (at z0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 691 | by eventually_elim (use X True in fastforce) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 692 | hence "frequently (\<lambda>w. w \<in> g -` pts \<inter> B) (at z0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 693 | by (meson at_neq_bot eventually_frequently) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 694 | thus "z0 \<notin> isolated_points_of (g -` pts \<inter> B)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 695 | unfolding isolated_points_of_def by (auto simp: frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 696 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 697 | moreover have "eventually (\<lambda>x. x \<in> X) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 698 | by (intro eventually_at_in_open') fact+ | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 699 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 700 | by (auto elim!: eventually_mono) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 701 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 702 | thus "\<not>z islimpt isolated_points_of (g -` pts \<inter> B)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 703 | by (auto simp: islimpt_conv_frequently_at frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 704 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 705 |   have "f \<circ> g analytic_on (\<Union>z\<in>B - isolated_points_of (g -` pts \<inter> B). {z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 706 | unfolding analytic_on_UN | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 707 | proof | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 708 | fix z assume z: "z \<in> B - isolated_points_of (g -` pts \<inter> B)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 709 | hence "z \<in> B" by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 710 |     have g': "g analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 711 | by (rule holomorphic_on_imp_analytic_at[OF g]) (use assms z in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 712 |     show "f \<circ> g analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 713 | proof (cases "g z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 714 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 715 | show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 716 | proof (rule analytic_on_compose) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 717 |         show "f analytic_on g ` {z}" using False z assms
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 718 | by (auto intro!: meromorphic_on_imp_analytic_at[OF f]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 719 | qed fact | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 720 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 721 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 722 | show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 723 | proof (cases "isolated_zero (\<lambda>w. g w - g z) z") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 724 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 725 | hence "eventually (\<lambda>w. g w - g z = 0) (nhds z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 726 | by (rule non_isolated_zero) (auto intro!: analytic_intros g') | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 727 |         hence "f \<circ> g analytic_on {z} \<longleftrightarrow> (\<lambda>_. f (g z)) analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 728 | by (intro analytic_at_cong) (auto elim!: eventually_mono) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 729 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 730 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 731 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 732 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 733 | hence ev: "eventually (\<lambda>w. g w \<noteq> g z) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 734 | by (auto simp: isolated_zero_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 735 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 736 | have "\<not>g z islimpt pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 737 | using \<open>g z \<in> pts\<close> f by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 738 | hence "eventually (\<lambda>w. w \<notin> pts) (at (g z))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 739 | by (auto simp: islimpt_conv_frequently_at frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 740 | moreover have "g \<midarrow>z\<rightarrow> g z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 741 | using analytic_at_imp_isCont[OF g'] isContD by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 742 | with ev have "filterlim g (at (g z)) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 743 | by (auto simp: filterlim_at) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 744 | ultimately have "eventually (\<lambda>w. g w \<notin> pts) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 745 | using eventually_compose_filterlim by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 746 | hence "z \<in> isolated_points_of (g -` pts \<inter> B)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 747 | using \<open>g z \<in> pts\<close> \<open>z \<in> B\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 748 | by (auto simp: isolated_points_of_def elim!: eventually_mono) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 749 | with z show ?thesis by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 750 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 751 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 752 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 753 | also have "\<dots> = B - isolated_points_of (g -` pts \<inter> B)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 754 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 755 | finally show "(\<lambda>x. f (g x)) holomorphic_on B - isolated_points_of (g -` pts \<inter> B)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 756 | unfolding o_def using analytic_imp_holomorphic by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 757 | qed (auto simp: isolated_points_of_def \<open>open B\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 758 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 759 | lemma meromorphic_on_compose': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 760 | assumes f: "f meromorphic_on A pts" and g: "g holomorphic_on B" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 761 | assumes "open B" and "g ` B \<subseteq> A" and "pts' = (isolated_points_of (g -` pts \<inter> B))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 762 | shows "(\<lambda>x. f (g x)) meromorphic_on B pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 763 | using meromorphic_on_compose[OF assms(1-4)] assms(5) by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 764 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 765 | lemma meromorphic_on_inverse': "inverse meromorphic_on UNIV 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 766 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 767 | by (auto intro!: holomorphic_intros singularity_intros not_essential_inverse | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 768 | isolated_singularity_at_inverse simp: islimpt_finite) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 769 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 770 | lemma meromorphic_on_inverse [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 771 | assumes mero: "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 772 |   shows   "(\<lambda>z. inverse (f z)) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 773 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 774 | have "open A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 775 | using mero by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 776 | have open': "open (A - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 777 | by (intro meromorphic_imp_open_diff[OF mero]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 778 | have holo: "f holomorphic_on A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 779 | using assms by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 780 | have ana: "f analytic_on A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 781 | using open' holo by (simp add: analytic_on_open) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 782 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 783 | show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 784 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 785 | proof (intro conjI ballI) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 786 |     fix z assume z: "z \<in> pts \<union> {z\<in>A. isolated_zero f z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 787 | have "isolated_singularity_at f z \<and> not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 788 | proof (cases "z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 789 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 790 |       have "f holomorphic_on A - pts - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 791 | by (intro holomorphic_on_subset[OF holo]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 792 | hence "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 793 | by (rule isolated_singularity_at_holomorphic) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 794 | (use z False in \<open>auto intro!: meromorphic_imp_open_diff[OF mero]\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 795 | moreover have "not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 796 | using z False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 797 | by (intro not_essential_holomorphic[OF holo] meromorphic_imp_open_diff[OF mero]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 798 | ultimately show ?thesis by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 799 | qed (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 800 | thus "isolated_singularity_at (\<lambda>z. inverse (f z)) z" "not_essential (\<lambda>z. inverse (f z)) z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 801 | by (auto intro!: isolated_singularity_at_inverse not_essential_inverse) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 802 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 803 | fix z assume "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 804 |     hence "\<not> z islimpt {z\<in>A. isolated_zero f z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 805 | by (rule not_islimpt_isolated_zeros[OF mero]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 806 |     thus "\<not> z islimpt pts \<union> {z \<in> A. isolated_zero f z}" using \<open>z \<in> A\<close>
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 807 | using mero by (auto simp: islimpt_Un meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 808 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 809 |     show "pts \<union> {z \<in> A. isolated_zero f z} \<subseteq> A"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 810 | using mero by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 811 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 812 |     have "(\<lambda>z. inverse (f z)) analytic_on (\<Union>w\<in>A - (pts \<union> {z \<in> A. isolated_zero f z}) . {w})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 813 | unfolding analytic_on_UN | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 814 | proof (intro ballI) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 815 |       fix w assume w: "w \<in> A - (pts \<union> {z \<in> A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 816 |       show "(\<lambda>z. inverse (f z)) analytic_on {w}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 817 | proof (cases "f w = 0") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 818 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 819 | thus ?thesis using w | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 820 | by (intro analytic_intros analytic_on_subset[OF ana]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 821 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 822 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 823 | have "eventually (\<lambda>w. f w = 0) (nhds w)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 824 | using True w by (intro non_isolated_zero analytic_on_subset[OF ana]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 825 |         hence "(\<lambda>z. inverse (f z)) analytic_on {w} \<longleftrightarrow> (\<lambda>_. 0) analytic_on {w}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 826 | using w by (intro analytic_at_cong refl) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 827 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 828 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 829 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 830 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 831 |     also have "\<dots> = A - (pts \<union> {z \<in> A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 832 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 833 | finally have "(\<lambda>z. inverse (f z)) analytic_on \<dots>" . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 834 |     moreover have "open (A - (pts \<union> {z \<in> A. isolated_zero f z}))"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 835 | using closedin_isolated_zeros[OF mero] open' \<open>open A\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 836 | by (metis (no_types, lifting) Diff_Diff_Int Diff_Un closedin_closed open_Diff open_Int) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 837 |     ultimately show "(\<lambda>z. inverse (f z)) holomorphic_on A - (pts \<union> {z \<in> A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 838 | by (subst (asm) analytic_on_open) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 839 | qed (use assms in \<open>auto simp: meromorphic_on_def islimpt_Un | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 840 | intro!: holomorphic_intros singularity_intros\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 841 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 842 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 843 | lemma meromorphic_on_inverse'' [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 844 |   assumes "f meromorphic_on A pts" "{z\<in>A. f z = 0} \<subseteq> pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 845 | shows "(\<lambda>z. inverse (f z)) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 846 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 847 |   have "(\<lambda>z. inverse (f z)) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 848 | by (intro meromorphic_on_inverse assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 849 |   also have "(pts \<union> {z \<in> A. isolated_zero f z}) = pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 850 | using assms(2) by (auto simp: isolated_zero_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 851 | finally show ?thesis . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 852 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 853 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 854 | lemma meromorphic_on_divide [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 855 | assumes "f meromorphic_on A pts" and "g meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 856 |   shows   "(\<lambda>z. f z / g z) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero g z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 857 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 858 |   have mero1: "(\<lambda>z. inverse (g z)) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero g z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 859 | by (intro meromorphic_intros assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 860 |   have sparse: "\<forall>x\<in>A. \<not> x islimpt pts \<union> {z\<in>A. isolated_zero g z}" and "pts \<subseteq> A"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 861 | using mero1 by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 862 |   have mero2: "f meromorphic_on A (pts \<union> {z\<in>A. isolated_zero g z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 863 | by (rule meromorphic_on_superset_pts[OF assms(1)]) (use sparse \<open>pts \<subseteq> A\<close> in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 864 |   have "(\<lambda>z. f z * inverse (g z)) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero g z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 865 | by (intro meromorphic_on_mult mero1 mero2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 866 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 867 | by (simp add: field_simps) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 868 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 869 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 870 | lemma meromorphic_on_divide' [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 871 |   assumes "f meromorphic_on A pts" "g meromorphic_on A pts" "{z\<in>A. g z = 0} \<subseteq> pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 872 | shows "(\<lambda>z. f z / g z) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 873 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 874 | have "(\<lambda>z. f z * inverse (g z)) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 875 | by (intro meromorphic_intros assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 876 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 877 | by (simp add: field_simps) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 878 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 879 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 880 | lemma meromorphic_on_cmult_left [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 881 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 882 | shows "(\<lambda>x. c * f x) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 883 | using assms by (intro meromorphic_intros) (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 884 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 885 | lemma meromorphic_on_cmult_right [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 886 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 887 | shows "(\<lambda>x. f x * c) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 888 | using assms by (intro meromorphic_intros) (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 889 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 890 | lemma meromorphic_on_scaleR [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 891 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 892 | shows "(\<lambda>x. c *\<^sub>R f x) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 893 | using assms unfolding scaleR_conv_of_real | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 894 | by (intro meromorphic_intros) (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 895 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 896 | lemma meromorphic_on_sum [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 897 | assumes "\<And>y. y \<in> I \<Longrightarrow> f y meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 898 |   assumes "I \<noteq> {} \<or> open A \<and> pts \<subseteq> A \<and> (\<forall>x\<in>A. \<not>x islimpt pts)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 899 | shows "(\<lambda>x. \<Sum>y\<in>I. f y x) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 900 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 901 | have *: "open A \<and> pts \<subseteq> A \<and> (\<forall>x\<in>A. \<not>x islimpt pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 902 | using assms(2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 903 | proof | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 904 |     assume "I \<noteq> {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 905 | then obtain x where "x \<in> I" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 906 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 907 | from assms(1)[OF this] show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 908 | by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 909 | qed auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 910 | show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 911 | using assms(1) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 912 | by (induction I rule: infinite_finite_induct) (use * in \<open>auto intro!: meromorphic_intros\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 913 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 914 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 915 | lemma meromorphic_on_prod [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 916 | assumes "\<And>y. y \<in> I \<Longrightarrow> f y meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 917 |   assumes "I \<noteq> {} \<or> open A \<and> pts \<subseteq> A \<and> (\<forall>x\<in>A. \<not>x islimpt pts)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 918 | shows "(\<lambda>x. \<Prod>y\<in>I. f y x) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 919 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 920 | have *: "open A \<and> pts \<subseteq> A \<and> (\<forall>x\<in>A. \<not>x islimpt pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 921 | using assms(2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 922 | proof | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 923 |     assume "I \<noteq> {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 924 | then obtain x where "x \<in> I" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 925 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 926 | from assms(1)[OF this] show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 927 | by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 928 | qed auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 929 | show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 930 | using assms(1) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 931 | by (induction I rule: infinite_finite_induct) (use * in \<open>auto intro!: meromorphic_intros\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 932 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 933 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 934 | lemma meromorphic_on_power [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 935 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 936 | shows "(\<lambda>x. f x ^ n) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 937 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 938 |   have "(\<lambda>x. \<Prod>i\<in>{..<n}. f x) meromorphic_on A pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 939 | by (intro meromorphic_intros assms(1)) (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 940 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 941 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 942 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 943 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 944 | lemma meromorphic_on_power_int [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 945 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 946 |   shows   "(\<lambda>z. f z powi n) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 947 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 948 |   have inv: "(\<lambda>x. inverse (f x)) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 949 | by (intro meromorphic_intros assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 950 |   have *: "f meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 951 | by (intro meromorphic_on_superset_pts [OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 952 | (use inv in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 953 | show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 954 | proof (cases "n \<ge> 0") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 955 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 956 |     have "(\<lambda>x. f x ^ nat n) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 957 | by (intro meromorphic_intros *) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 958 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 959 | using True by (simp add: power_int_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 960 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 961 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 962 |     have "(\<lambda>x. inverse (f x) ^ nat (-n)) meromorphic_on A (pts \<union> {z \<in> A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 963 | by (intro meromorphic_intros assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 964 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 965 | using False by (simp add: power_int_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 966 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 967 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 968 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 969 | lemma meromorphic_on_power_int' [meromorphic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 970 | assumes "f meromorphic_on A pts" "n \<ge> 0 \<or> (\<forall>z\<in>A. isolated_zero f z \<longrightarrow> z \<in> pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 971 | shows "(\<lambda>z. f z powi n) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 972 | proof (cases "n \<ge> 0") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 973 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 974 | have "(\<lambda>z. f z ^ nat n) meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 975 | by (intro meromorphic_intros assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 976 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 977 | using True by (simp add: power_int_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 978 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 979 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 980 |   have "(\<lambda>z. f z powi n) meromorphic_on A (pts \<union> {z\<in>A. isolated_zero f z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 981 | by (rule meromorphic_on_power_int) fact | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 982 |   also from assms(2) False have "pts \<union> {z\<in>A. isolated_zero f z} = pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 983 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 984 | finally show ?thesis . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 985 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 986 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 987 | lemma has_laurent_expansion_on_imp_meromorphic_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 988 | assumes "open A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 989 | assumes laurent: "\<And>z. z \<in> A \<Longrightarrow> \<exists>F. (\<lambda>w. f (z + w)) has_laurent_expansion F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 990 |   shows   "f meromorphic_on A {z\<in>A. \<not>f analytic_on {z}}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 991 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 992 | proof (intro conjI ballI) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 993 |   fix z assume "z \<in> {z\<in>A. \<not>f analytic_on {z}}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 994 | then obtain F where F: "(\<lambda>w. f (z + w)) has_laurent_expansion F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 995 | using laurent[of z] by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 996 | from F show "not_essential f z" "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 997 | using has_laurent_expansion_not_essential has_laurent_expansion_isolated by blast+ | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 998 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 999 | fix z assume z: "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1000 | obtain F where F: "(\<lambda>w. f (z + w)) has_laurent_expansion F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1001 | using laurent[of z] \<open>z \<in> A\<close> by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1002 | from F have "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1003 | using has_laurent_expansion_isolated z by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1004 |   then obtain r where r: "r > 0" "f analytic_on ball z r - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1005 | unfolding isolated_singularity_at_def by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1006 |   have "f analytic_on {w}" if "w \<in> ball z r - {z}" for w
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1007 | by (rule analytic_on_subset[OF r(2)]) (use that in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1008 |   hence "eventually (\<lambda>w. f analytic_on {w}) (at z)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1009 | using eventually_at_in_open[of "ball z r" z] \<open>r > 0\<close> by (auto elim!: eventually_mono) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1010 |   hence "\<not>z islimpt {w. \<not>f analytic_on {w}}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1011 | by (auto simp: islimpt_conv_frequently_at frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1012 |   thus "\<not>z islimpt {w\<in>A. \<not>f analytic_on {w}}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1013 |     using islimpt_subset[of z "{w\<in>A. \<not>f analytic_on {w}}" "{w. \<not>f analytic_on {w}}"] by blast
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1014 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1015 |   have "f analytic_on A - {w\<in>A. \<not>f analytic_on {w}}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1016 | by (subst analytic_on_analytic_at) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1017 |   thus "f holomorphic_on A - {w\<in>A. \<not>f analytic_on {w}}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1018 | by (meson analytic_imp_holomorphic) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1019 | qed (use assms in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1020 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1021 | lemma meromorphic_on_imp_has_laurent_expansion: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1022 | assumes "f meromorphic_on A pts" "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1023 | shows "(\<lambda>w. f (z + w)) has_laurent_expansion laurent_expansion f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1024 | proof (cases "z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1025 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1026 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1027 | using assms by (intro not_essential_has_laurent_expansion) (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1028 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1029 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1030 | have "f holomorphic_on (A - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1031 | using assms by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1032 | moreover have "z \<in> A - pts" "open (A - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1033 | using assms(2) False by (auto intro!: meromorphic_imp_open_diff[OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1034 |   ultimately have "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1035 | unfolding analytic_at by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1036 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1037 | using isolated_singularity_at_analytic not_essential_analytic | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1038 | not_essential_has_laurent_expansion by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1039 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1040 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1041 | lemma | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1042 | assumes "isolated_singularity_at f z" "f \<midarrow>z\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1043 | shows eventually_remove_sings_eq_nhds': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1044 | "eventually (\<lambda>w. remove_sings f w = (if w = z then c else f w)) (nhds z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1045 |     and   remove_sings_analytic_at_singularity: "remove_sings f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1046 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1047 | have "eventually (\<lambda>w. w \<noteq> z) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1048 | by (auto simp: eventually_at_filter) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1049 | hence "eventually (\<lambda>w. remove_sings f w = (if w = z then c else f w)) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1050 | using eventually_remove_sings_eq_at[OF assms(1)] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1051 | by eventually_elim auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1052 | moreover have "remove_sings f z = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1053 | using assms by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1054 | ultimately show ev: "eventually (\<lambda>w. remove_sings f w = (if w = z then c else f w)) (nhds z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1055 | by (simp add: eventually_at_filter) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1056 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1057 |   have "(\<lambda>w. if w = z then c else f w) analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1058 | by (intro removable_singularity' assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1059 |   also have "?this \<longleftrightarrow> remove_sings f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1060 | using ev by (intro analytic_at_cong) (auto simp: eq_commute) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1061 | finally show \<dots> . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1062 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1063 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1064 | lemma remove_sings_meromorphic_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1065 | assumes "f meromorphic_on A pts" "\<And>z. z \<in> pts - pts' \<Longrightarrow> \<not>is_pole f z" "pts' \<subseteq> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1066 | shows "remove_sings f meromorphic_on A pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1067 | unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1068 | proof safe | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1069 |   have "remove_sings f analytic_on {z}" if "z \<in> A - pts'" for z
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1070 | proof (cases "z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1071 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1072 |     hence *: "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1073 | using assms meromorphic_imp_open_diff[OF assms(1)] that | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1074 | by (force simp: meromorphic_on_def analytic_at) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1075 |     have "remove_sings f analytic_on {z} \<longleftrightarrow> f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1076 | by (intro analytic_at_cong eventually_remove_sings_eq_nhds * refl) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1077 | thus ?thesis using * by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1078 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1079 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1080 | have isol: "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1081 | using True using assms by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1082 | from assms(1) have "not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1083 | using True by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1084 | with assms(2) True that obtain c where "f \<midarrow>z\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1085 | by (auto simp: not_essential_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1086 |     thus "remove_sings f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1087 | by (intro remove_sings_analytic_at_singularity isol) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1088 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1089 | hence "remove_sings f analytic_on A - pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1090 | by (subst analytic_on_analytic_at) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1091 | thus "remove_sings f holomorphic_on A - pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1092 | using meromorphic_imp_open_diff'[OF assms(1,3)] by (subst (asm) analytic_on_open) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1093 | qed (use assms islimpt_subset[OF _ assms(3)] in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1094 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1095 | lemma remove_sings_holomorphic_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1096 | assumes "f meromorphic_on A pts" "\<And>z. z \<in> pts \<Longrightarrow> \<not>is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1097 | shows "remove_sings f holomorphic_on A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1098 |   using remove_sings_meromorphic_on[OF assms(1), of "{}"] assms(2)
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1099 | by (auto simp: meromorphic_on_no_singularities) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1100 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1101 | lemma meromorphic_on_Ex_iff: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1102 | "(\<exists>pts. f meromorphic_on A pts) \<longleftrightarrow> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1103 | open A \<and> (\<forall>z\<in>A. \<exists>F. (\<lambda>w. f (z + w)) has_laurent_expansion F)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1104 | proof safe | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1105 | fix pts assume *: "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1106 | from * show "open A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1107 | by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1108 | show "\<exists>F. (\<lambda>w. f (z + w)) has_laurent_expansion F" if "z \<in> A" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1109 | using that * | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1110 | by (intro exI[of _ "laurent_expansion f z"] meromorphic_on_imp_has_laurent_expansion) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1111 | qed (blast intro!: has_laurent_expansion_on_imp_meromorphic_on) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1112 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1113 | lemma is_pole_inverse_holomorphic_pts: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1114 | fixes pts::"complex set" and f::"complex \<Rightarrow> complex" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1115 | defines "g \<equiv> \<lambda>x. (if x\<in>pts then 0 else inverse (f x))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1116 | assumes mer: "f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1117 | and non_z: "\<And>z. z \<in> D - pts \<Longrightarrow> f z \<noteq> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1118 | and all_poles:"\<forall>x. is_pole f x \<longleftrightarrow> x\<in>pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1119 | shows "g holomorphic_on D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1120 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1121 | have "open D" and f_holo: "f holomorphic_on (D-pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1122 | using mer by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1123 |   have "\<exists>r. r>0 \<and> f analytic_on ball z r - {z} 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1124 |             \<and> (\<forall>x \<in> ball z r - {z}. f x\<noteq>0)" if "z\<in>pts" for z 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1125 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1126 | have "isolated_singularity_at f z" "is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1127 | using mer meromorphic_on_def that all_poles by blast+ | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1128 |     then obtain r1 where "r1>0" and fan: "f analytic_on ball z r1 - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1129 | by (meson isolated_singularity_at_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1130 |     obtain r2 where "r2>0" "\<forall>x \<in> ball z r2 - {z}. f x\<noteq>0"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1131 | using non_zero_neighbour_pole[OF \<open>is_pole f z\<close>] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1132 | unfolding eventually_at by (metis Diff_iff UNIV_I dist_commute insertI1 mem_ball) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1133 | define r where "r = min r1 r2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1134 | have "r>0" by (simp add: \<open>0 < r2\<close> \<open>r1>0\<close> r_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1135 |     moreover have "f analytic_on ball z r - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1136 | using r_def by (force intro: analytic_on_subset [OF fan]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1137 |     moreover have "\<forall>x \<in> ball z r - {z}. f x\<noteq>0"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1138 |       by (simp add: \<open>\<forall>x\<in>ball z r2 - {z}. f x \<noteq> 0\<close> r_def)
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1139 | ultimately show ?thesis by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1140 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1141 | then obtain get_r where r_pos:"get_r z>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1142 |       and r_ana:"f analytic_on ball z (get_r z) - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1143 |       and r_nz:"\<forall>x \<in> ball z (get_r z) - {z}. f x\<noteq>0"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1144 | if "z\<in>pts" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1145 | by metis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1146 | define p_balls where "p_balls \<equiv> \<Union>z\<in>pts. ball z (get_r z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1147 | have g_ball:"g holomorphic_on ball z (get_r z)" if "z\<in>pts" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1148 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1149 | have "(\<lambda>x. if x = z then 0 else inverse (f x)) holomorphic_on ball z (get_r z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1150 | proof (rule is_pole_inverse_holomorphic) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1151 |       show "f holomorphic_on ball z (get_r z) - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1152 | using analytic_imp_holomorphic r_ana that by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1153 | show "is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1154 | using mer meromorphic_on_def that all_poles by force | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1155 |       show "\<forall>x\<in>ball z (get_r z) - {z}. f x \<noteq> 0"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1156 | using r_nz that by metis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1157 | qed auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1158 | then show ?thesis unfolding g_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1159 | by (smt (verit, ccfv_SIG) Diff_iff Elementary_Metric_Spaces.open_ball | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1160 | all_poles analytic_imp_holomorphic empty_iff | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1161 | holomorphic_transform insert_iff not_is_pole_holomorphic | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1162 | open_delete r_ana that) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1163 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1164 | then have "g holomorphic_on p_balls" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1165 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1166 | have "g analytic_on p_balls" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1167 | unfolding p_balls_def analytic_on_UN | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1168 | using g_ball by (simp add: analytic_on_open) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1169 | moreover have "open p_balls" using p_balls_def by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1170 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1171 | by (simp add: analytic_imp_holomorphic) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1172 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1173 | moreover have "g holomorphic_on D-pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1174 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1175 | have "(\<lambda>z. inverse (f z)) holomorphic_on D - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1176 | using f_holo holomorphic_on_inverse non_z by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1177 | then show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1178 | by (metis DiffD2 g_def holomorphic_transform) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1179 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1180 | moreover have "open p_balls" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1181 | using p_balls_def by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1182 | ultimately have "g holomorphic_on (p_balls \<union> (D-pts))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1183 | by (simp add: holomorphic_on_Un meromorphic_imp_open_diff[OF mer]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1184 | moreover have "D \<subseteq> p_balls \<union> (D-pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1185 | unfolding p_balls_def using \<open>\<And>z. z \<in> pts \<Longrightarrow> 0 < get_r z\<close> by force | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1186 | ultimately show "g holomorphic_on D" by (meson holomorphic_on_subset) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1187 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1188 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1189 | lemma meromorphic_imp_analytic_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1190 | assumes "f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1191 | shows "f analytic_on (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1192 | by (metis assms analytic_on_open meromorphic_imp_open_diff meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1193 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1194 | lemma meromorphic_imp_constant_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1195 | assumes merf: "f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1196 | and "f constant_on (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1197 | and "\<forall>x\<in>pts. is_pole f x" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1198 | shows "f constant_on D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1199 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1200 | obtain c where c:"\<And>z. z \<in> D-pts \<Longrightarrow> f z = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1201 | by (meson assms constant_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1202 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1203 | have "f z = c" if "z \<in> D" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1204 | proof (cases "is_pole f z") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1205 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1206 |     then obtain r0 where "r0 > 0" and r0: "f analytic_on ball z r0 - {z}" and pol: "is_pole f z"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1207 | using merf unfolding meromorphic_on_def isolated_singularity_at_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1208 | by (metis \<open>z \<in> D\<close> insert_Diff insert_Diff_if insert_iff merf | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1209 | meromorphic_imp_open_diff not_is_pole_holomorphic) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1210 | have "open D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1211 | using merf meromorphic_on_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1212 | then obtain r where "r > 0" "ball z r \<subseteq> D" "r \<le> r0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1213 | by (smt (verit, best) \<open>0 < r0\<close> \<open>z \<in> D\<close> openE order_subst2 subset_ball) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1214 |     have r: "f analytic_on ball z r - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1215 | by (meson Diff_mono \<open>r \<le> r0\<close> analytic_on_subset order_refl r0 subset_ball) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1216 |     have "ball z r - {z} \<subseteq> -pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1217 | using merf r unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1218 | by (meson ComplI Elementary_Metric_Spaces.open_ball | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1219 | analytic_imp_holomorphic assms(3) not_is_pole_holomorphic open_delete subsetI) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1220 |     with \<open>ball z r \<subseteq> D\<close> have "ball z r - {z} \<subseteq> D-pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1221 | by fastforce | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1222 |     with c have c': "\<And>u. u \<in> ball z r - {z} \<Longrightarrow> f u = c"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1223 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1224 | have False if "\<forall>\<^sub>F x in at z. cmod c + 1 \<le> cmod (f x)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1225 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1226 |       have "\<forall>\<^sub>F x in at z within ball z r - {z}. cmod c + 1 \<le> cmod (f x)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1227 | by (smt (verit, best) Diff_UNIV Diff_eq_empty_iff eventually_at_topological insert_subset that) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1228 | with \<open>r > 0\<close> show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1229 | apply (simp add: c' eventually_at_filter topological_space_class.eventually_nhds open_dist) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1230 | by (metis dist_commute min_less_iff_conj perfect_choose_dist) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1231 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1232 | with pol show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1233 | by (auto simp: is_pole_def filterlim_at_infinity_conv_norm_at_top filterlim_at_top) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1234 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1235 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1236 | then show ?thesis by (meson DiffI assms(3) c that) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1237 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1238 | then show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1239 | by (simp add: constant_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1240 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1241 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1242 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1243 | lemma meromorphic_isolated: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1244 | assumes merf: "f meromorphic_on D pts" and "p\<in>pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1245 |   obtains r where "r>0" "ball p r \<subseteq> D" "ball p r \<inter> pts = {p}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1246 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1247 | have "\<forall>z\<in>D. \<exists>e>0. finite (pts \<inter> ball z e)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1248 | using merf unfolding meromorphic_on_def islimpt_eq_infinite_ball | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1249 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1250 | then obtain r0 where r0:"r0>0" "finite (pts \<inter> ball p r0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1251 | by (metis assms(2) in_mono merf meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1252 |   moreover define pts' where "pts' = pts \<inter> ball p r0 - {p}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1253 | ultimately have "finite pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1254 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1255 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1256 |   define r1 where "r1=(if pts'={} then r0 else 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1257 |                           min (Min {dist p' p |p'. p'\<in>pts'}/2) r0)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1258 |   have "r1>0 \<and> pts \<inter> ball p r1 - {p} = {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1259 |   proof (cases "pts'={}")
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1260 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1261 | then show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1262 | using pts'_def r0(1) r1_def by presburger | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1263 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1264 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1265 |     define S where "S={dist p' p |p'. p'\<in>pts'}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1266 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1267 |     have nempty:"S \<noteq> {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1268 | using False S_def by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1269 | have finite:"finite S" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1270 | using \<open>finite pts'\<close> S_def by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1271 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1272 | have "r1>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1273 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1274 | have "r1=min (Min S/2) r0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1275 | using False unfolding S_def r1_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1276 | moreover have "Min S\<in>S" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1277 |         using \<open>S\<noteq>{}\<close> \<open>finite S\<close>  Min_in by auto
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1278 | then have "Min S>0" unfolding S_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1279 | using pts'_def by force | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1280 | ultimately show ?thesis using \<open>r0>0\<close> by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1281 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1282 |     moreover have "pts \<inter> ball p r1 - {p} = {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1283 | proof (rule ccontr) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1284 |       assume "pts \<inter> ball p r1 - {p} \<noteq> {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1285 |       then obtain p' where "p'\<in>pts \<inter> ball p r1 - {p}" by blast
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1286 | moreover have "r1\<le>r0" using r1_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1287 | ultimately have "p'\<in>pts'" unfolding pts'_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1288 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1289 | then have "dist p' p\<ge>Min S" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1290 | using S_def eq_Min_iff local.finite by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1291 | moreover have "dist p' p < Min S" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1292 |         using \<open>p'\<in>pts \<inter> ball p r1 - {p}\<close> False unfolding r1_def
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1293 | apply (fold S_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1294 | by (smt (verit, ccfv_threshold) DiffD1 Int_iff dist_commute | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1295 | dist_triangle_half_l mem_ball) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1296 | ultimately show False by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1297 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1298 | ultimately show ?thesis by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1299 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1300 |   then have "r1>0" and r1_pts:"pts \<inter> ball p r1 - {p} = {}" by auto
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1301 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1302 | obtain r2 where "r2>0" "ball p r2 \<subseteq> D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1303 | by (metis assms(2) merf meromorphic_on_def openE subset_eq) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1304 | define r where "r=min r1 r2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1305 | have "r > 0" unfolding r_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1306 | by (simp add: \<open>0 < r1\<close> \<open>0 < r2\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1307 | moreover have "ball p r \<subseteq> D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1308 | using \<open>ball p r2 \<subseteq> D\<close> r_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1309 |   moreover have "ball p r \<inter> pts = {p}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1310 | using assms(2) \<open>r>0\<close> r1_pts | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1311 | unfolding r_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1312 | ultimately show ?thesis using that by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1313 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1314 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1315 | lemma meromorphic_pts_closure: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1316 | assumes merf: "f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1317 | shows "pts \<subseteq> closure (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1318 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1319 | have "p islimpt (D - pts)" if "p\<in>pts" for p | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1320 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1321 |     obtain r where "r>0" "ball p r \<subseteq> D" "ball p r \<inter> pts = {p}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1322 | using meromorphic_isolated[OF merf \<open>p\<in>pts\<close>] by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1323 | from \<open>r>0\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1324 |     have "p islimpt ball p r - {p}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1325 | by (meson open_ball ball_subset_cball in_mono islimpt_ball | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1326 | islimpt_punctured le_less open_contains_ball_eq) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1327 |     moreover have " ball p r - {p} \<subseteq> D - pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1328 |       using \<open>ball p r \<inter> pts = {p}\<close> \<open>ball p r \<subseteq> D\<close> by fastforce
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1329 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1330 | using islimpt_subset by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1331 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1332 | then show ?thesis by (simp add: islimpt_in_closure subset_eq) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1333 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1334 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1335 | lemma nconst_imp_nzero_neighbour: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1336 | assumes merf: "f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1337 | and f_nconst:"\<not>(\<forall>w\<in>D-pts. f w=0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1338 | and "z\<in>D" and "connected D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1339 | shows "(\<forall>\<^sub>F w in at z. f w \<noteq> 0 \<and> w \<in> D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1340 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1341 | obtain \<beta> where \<beta>:"\<beta> \<in> D - pts" "f \<beta>\<noteq>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1342 | using f_nconst by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1343 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1344 | have ?thesis if "z\<notin>pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1345 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1346 | have "\<forall>\<^sub>F w in at z. f w \<noteq> 0 \<and> w \<in> D - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1347 | apply (rule non_zero_neighbour_alt[of f "D-pts" z \<beta>]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1348 | subgoal using merf meromorphic_on_def by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1349 | subgoal using merf meromorphic_imp_open_diff by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1350 | subgoal using assms(4) merf meromorphic_imp_connected_diff by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1351 | subgoal by (simp add: assms(3) that) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1352 | using \<beta> by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1353 | then show ?thesis by (auto elim:eventually_mono) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1354 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1355 | moreover have ?thesis if "z\<in>pts" "\<not> f \<midarrow>z\<rightarrow> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1356 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1357 | have "\<forall>\<^sub>F w in at z. w \<in> D - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1358 | using merf[unfolded meromorphic_on_def islimpt_iff_eventually] \<open>z\<in>D\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1359 | using eventually_at_in_open' eventually_elim2 by fastforce | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1360 | moreover have "\<forall>\<^sub>F w in at z. f w \<noteq> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1361 | proof (cases "is_pole f z") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1362 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1363 | then show ?thesis using non_zero_neighbour_pole by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1364 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1365 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1366 | moreover have "not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1367 | using merf meromorphic_on_def that(1) by fastforce | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1368 | ultimately obtain c where "c\<noteq>0" "f \<midarrow>z\<rightarrow> c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1369 | by (metis \<open>\<not> f \<midarrow>z\<rightarrow> 0\<close> not_essential_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1370 | then show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1371 | using tendsto_imp_eventually_ne by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1372 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1373 | ultimately show ?thesis by eventually_elim auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1374 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1375 | moreover have ?thesis if "z\<in>pts" "f \<midarrow>z\<rightarrow> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1376 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1377 | define ff where "ff=(\<lambda>x. if x=z then 0 else f x)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1378 |     define A where "A=D - (pts - {z})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1379 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1380 |     have "f holomorphic_on A - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1381 | by (metis A_def Diff_insert analytic_imp_holomorphic | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1382 | insert_Diff merf meromorphic_imp_analytic_on that(1)) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1383 | moreover have "open A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1384 | using A_def merf meromorphic_imp_open_diff' by force | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1385 | ultimately have "ff holomorphic_on A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1386 | using \<open>f \<midarrow>z\<rightarrow> 0\<close> unfolding ff_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1387 | by (rule removable_singularity) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1388 | moreover have "connected A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1389 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1390 | have "connected (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1391 | using assms(4) merf meromorphic_imp_connected_diff by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1392 | moreover have "D - pts \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1393 | unfolding A_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1394 | moreover have "A \<subseteq> closure (D - pts)" unfolding A_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1395 | by (smt (verit, ccfv_SIG) Diff_empty Diff_insert | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1396 | closure_subset insert_Diff_single insert_absorb | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1397 | insert_subset merf meromorphic_pts_closure that(1)) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1398 | ultimately show ?thesis using connected_intermediate_closure | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1399 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1400 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1401 | moreover have "z \<in> A" using A_def assms(3) by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1402 | moreover have "ff z = 0" unfolding ff_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1403 | moreover have "\<beta> \<in> A " using A_def \<beta>(1) by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1404 | moreover have "ff \<beta> \<noteq> 0" using \<beta>(1) \<beta>(2) ff_def that(1) by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1405 | ultimately obtain r where "0 < r" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1406 |         "ball z r \<subseteq> A" "\<And>x. x \<in> ball z r - {z} \<Longrightarrow> ff x \<noteq> 0"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1407 | using \<open>open A\<close> isolated_zeros[of ff A z \<beta>] by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1408 | then show ?thesis unfolding eventually_at ff_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1409 | by (intro exI[of _ r]) (auto simp: A_def dist_commute ball_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1410 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1411 | ultimately show ?thesis by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1412 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1413 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1414 | lemma nconst_imp_nzero_neighbour': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1415 | assumes merf: "f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1416 | and f_nconst:"\<not>(\<forall>w\<in>D-pts. f w=0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1417 | and "z\<in>D" and "connected D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1418 | shows "\<forall>\<^sub>F w in at z. f w \<noteq> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1419 | using nconst_imp_nzero_neighbour[OF assms] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1420 | by (auto elim:eventually_mono) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1421 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1422 | lemma meromorphic_compact_finite_zeros: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1423 | assumes merf:"f meromorphic_on D pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1424 | and "compact S" "S \<subseteq> D" "connected D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1425 | and f_nconst:"\<not>(\<forall>w\<in>D-pts. f w=0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1426 |   shows "finite ({x\<in>S. f x=0})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1427 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1428 |   have "finite ({x\<in>S. f x=0 \<and> x \<notin> pts})" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1429 | proof (rule ccontr) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1430 |     assume "infinite {x \<in> S. f x = 0 \<and> x \<notin> pts}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1431 |     then obtain z where "z\<in>S" and z_lim:"z islimpt {x \<in> S. f x = 0
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1432 | \<and> x \<notin> pts}" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1433 | using \<open>compact S\<close> unfolding compact_eq_Bolzano_Weierstrass | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1434 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1435 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1436 | from z_lim | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1437 | have "\<exists>\<^sub>F x in at z. f x = 0 \<and> x \<in> S \<and> x \<notin> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1438 | unfolding islimpt_iff_eventually not_eventually by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1439 | moreover have "\<forall>\<^sub>F w in at z. f w \<noteq> 0 \<and> w \<in> D - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1440 | using nconst_imp_nzero_neighbour[OF merf f_nconst _ \<open>connected D\<close>] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1441 | \<open>z\<in>S\<close> \<open>S \<subseteq> D\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1442 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1443 | ultimately have "\<exists>\<^sub>F x in at z. False" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1444 | by (simp add: eventually_mono frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1445 | then show False by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1446 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1447 | moreover have "finite (S \<inter> pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1448 | using meromorphic_compact_finite_pts[OF merf \<open>compact S\<close> \<open>S \<subseteq> D\<close>] . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1449 |   ultimately have "finite ({x\<in>S. f x=0 \<and> x \<notin> pts} \<union> (S \<inter> pts))"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1450 | unfolding finite_Un by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1451 | then show ?thesis by (elim rev_finite_subset) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1452 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1453 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1454 | lemma meromorphic_onI [intro?]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1455 | assumes "open A" "pts \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1456 | assumes "f holomorphic_on A - pts" "\<And>z. z \<in> A \<Longrightarrow> \<not>z islimpt pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1457 | assumes "\<And>z. z \<in> pts \<Longrightarrow> isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1458 | assumes "\<And>z. z \<in> pts \<Longrightarrow> not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1459 | shows "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1460 | using assms unfolding meromorphic_on_def by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1461 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1462 | lemma Polygamma_plus_of_nat: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1463 | assumes "\<forall>k<m. z \<noteq> -of_nat k" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1464 | shows "Polygamma n (z + of_nat m) = | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1465 | Polygamma n z + (-1) ^ n * fact n * (\<Sum>k<m. 1 / (z + of_nat k) ^ Suc n)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1466 | using assms | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1467 | proof (induction m) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1468 | case (Suc m) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1469 | have "Polygamma n (z + of_nat (Suc m)) = Polygamma n (z + of_nat m + 1)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1470 | by (simp add: add_ac) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1471 | also have "\<dots> = Polygamma n (z + of_nat m) + (-1) ^ n * fact n * (1 / ((z + of_nat m) ^ Suc n))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1472 | using Suc.prems by (subst Polygamma_plus1) (auto simp: add_eq_0_iff2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1473 | also have "Polygamma n (z + of_nat m) = | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1474 | Polygamma n z + (-1) ^ n * (\<Sum>k<m. 1 / (z + of_nat k) ^ Suc n) * fact n" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1475 | using Suc.prems by (subst Suc.IH) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1476 | finally show ?case | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1477 | by (simp add: algebra_simps) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1478 | qed auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1479 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1480 | lemma tendsto_Gamma [tendsto_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1481 | assumes "(f \<longlongrightarrow> c) F" "c \<notin> \<int>\<^sub>\<le>\<^sub>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1482 | shows "((\<lambda>z. Gamma (f z)) \<longlongrightarrow> Gamma c) F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1483 | by (intro isCont_tendsto_compose[OF _ assms(1)] continuous_intros assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1484 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1485 | lemma tendsto_Polygamma [tendsto_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1486 |   fixes f :: "_ \<Rightarrow> 'a :: {real_normed_field,euclidean_space}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1487 | assumes "(f \<longlongrightarrow> c) F" "c \<notin> \<int>\<^sub>\<le>\<^sub>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1488 | shows "((\<lambda>z. Polygamma n (f z)) \<longlongrightarrow> Polygamma n c) F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1489 | by (intro isCont_tendsto_compose[OF _ assms(1)] continuous_intros assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1490 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1491 | lemma analytic_on_Gamma' [analytic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1492 | assumes "f analytic_on A" "\<forall>x\<in>A. f x \<notin> \<int>\<^sub>\<le>\<^sub>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1493 | shows "(\<lambda>z. Gamma (f z)) analytic_on A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1494 | using analytic_on_compose_gen[OF assms(1) analytic_Gamma[of "f ` A"]] assms(2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1495 | by (auto simp: o_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1496 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1497 | lemma analytic_on_Polygamma' [analytic_intros]: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1498 | assumes "f analytic_on A" "\<forall>x\<in>A. f x \<notin> \<int>\<^sub>\<le>\<^sub>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1499 | shows "(\<lambda>z. Polygamma n (f z)) analytic_on A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1500 | using analytic_on_compose_gen[OF assms(1) analytic_on_Polygamma[of "f ` A" n]] assms(2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1501 | by (auto simp: o_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1502 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1503 | lemma | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1504 | shows is_pole_Polygamma: "is_pole (Polygamma n) (-of_nat m :: complex)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1505 | and zorder_Polygamma: "zorder (Polygamma n) (-of_nat m) = -int (Suc n)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1506 | and residue_Polygamma: "residue (Polygamma n) (-of_nat m) = (if n = 0 then -1 else 0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1507 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1508 | define g1 :: "complex \<Rightarrow> complex" where | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1509 | "g1 = (\<lambda>z. Polygamma n (z + of_nat (Suc m)) + | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1510 | (-1) ^ Suc n * fact n * (\<Sum>k<m. 1 / (z + of_nat k) ^ Suc n))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1511 | define g :: "complex \<Rightarrow> complex" where | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1512 | "g = (\<lambda>z. g1 z + (-1) ^ Suc n * fact n / (z + of_nat m) ^ Suc n)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1513 | define F where "F = fps_to_fls (fps_expansion g1 (-of_nat m)) + fls_const ((-1) ^ Suc n * fact n) / fls_X ^ Suc n" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1514 | have F_altdef: "F = fps_to_fls (fps_expansion g1 (-of_nat m)) + fls_shift (n+1) (fls_const ((-1) ^ Suc n * fact n))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1515 | by (simp add: F_def del: power_Suc) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1516 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1517 | have "\<not>(-of_nat m) islimpt (\<int>\<^sub>\<le>\<^sub>0 :: complex set)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1518 | by (intro discrete_imp_not_islimpt[where e = 1]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1519 | (auto elim!: nonpos_Ints_cases simp: dist_of_int) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1520 | hence "eventually (\<lambda>z::complex. z \<notin> \<int>\<^sub>\<le>\<^sub>0) (at (-of_nat m))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1521 | by (auto simp: islimpt_conv_frequently_at frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1522 | hence ev: "eventually (\<lambda>z. Polygamma n z = g z) (at (-of_nat m))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1523 | proof eventually_elim | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1524 | case (elim z) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1525 | hence *: "\<forall>k<Suc m. z \<noteq> - of_nat k" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1526 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1527 | thus ?case | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1528 | using Polygamma_plus_of_nat[of "Suc m" z n, OF *] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1529 | by (auto simp: g_def g1_def algebra_simps) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1530 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1531 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1532 | have "(\<lambda>w. g (-of_nat m + w)) has_laurent_expansion F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1533 | unfolding g_def F_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1534 | by (intro laurent_expansion_intros has_laurent_expansion_fps analytic_at_imp_has_fps_expansion) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1535 | (auto simp: g1_def intro!: laurent_expansion_intros analytic_intros) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1536 | also have "?this \<longleftrightarrow> (\<lambda>w. Polygamma n (-of_nat m + w)) has_laurent_expansion F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1537 | using ev by (intro has_laurent_expansion_cong refl) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1538 | (simp_all add: eq_commute at_to_0' eventually_filtermap) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1539 | finally have *: "(\<lambda>w. Polygamma n (-of_nat m + w)) has_laurent_expansion F" . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1540 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1541 | have subdegree: "fls_subdegree F = -int (Suc n)" unfolding F_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1542 | by (subst fls_subdegree_add_eq2) (simp_all add: fls_subdegree_fls_to_fps fls_divide_subdegree) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1543 | have [simp]: "F \<noteq> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1544 | using subdegree by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1545 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1546 | show "is_pole (Polygamma n) (-of_nat m :: complex)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1547 | using * by (rule has_laurent_expansion_imp_is_pole) (auto simp: subdegree) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1548 | show "zorder (Polygamma n) (-of_nat m :: complex) = -int (Suc n)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1549 | by (subst has_laurent_expansion_zorder[OF *]) (auto simp: subdegree) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1550 | show "residue (Polygamma n) (-of_nat m :: complex) = (if n = 0 then -1 else 0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1551 | by (subst has_laurent_expansion_residue[OF *]) (auto simp: F_altdef) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1552 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1553 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1554 | lemma Gamma_meromorphic_on [meromorphic_intros]: "Gamma meromorphic_on UNIV \<int>\<^sub>\<le>\<^sub>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1555 | proof | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1556 | show "\<not>z islimpt \<int>\<^sub>\<le>\<^sub>0" for z :: complex | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1557 | by (intro discrete_imp_not_islimpt[of 1]) (auto elim!: nonpos_Ints_cases simp: dist_of_int) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1558 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1559 | fix z :: complex assume z: "z \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1560 | then obtain n where n: "z = -of_nat n" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1561 | by (elim nonpos_Ints_cases') | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1562 | show "not_essential Gamma z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1563 | by (auto simp: n intro!: is_pole_imp_not_essential is_pole_Gamma) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1564 |   have *: "open (-(\<int>\<^sub>\<le>\<^sub>0 - {z}))"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1565 | by (intro open_Compl discrete_imp_closed[of 1]) (auto elim!: nonpos_Ints_cases simp: dist_of_int) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1566 |   have "Gamma holomorphic_on -(\<int>\<^sub>\<le>\<^sub>0 - {z}) - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1567 | by (intro holomorphic_intros) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1568 | thus "isolated_singularity_at Gamma z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1569 | by (rule isolated_singularity_at_holomorphic) (use z * in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1570 | qed (auto intro!: holomorphic_intros) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1571 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1572 | lemma Polygamma_meromorphic_on [meromorphic_intros]: "Polygamma n meromorphic_on UNIV \<int>\<^sub>\<le>\<^sub>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1573 | proof | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1574 | show "\<not>z islimpt \<int>\<^sub>\<le>\<^sub>0" for z :: complex | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1575 | by (intro discrete_imp_not_islimpt[of 1]) (auto elim!: nonpos_Ints_cases simp: dist_of_int) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1576 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1577 | fix z :: complex assume z: "z \<in> \<int>\<^sub>\<le>\<^sub>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1578 | then obtain m where n: "z = -of_nat m" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1579 | by (elim nonpos_Ints_cases') | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1580 | show "not_essential (Polygamma n) z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1581 | by (auto simp: n intro!: is_pole_imp_not_essential is_pole_Polygamma) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1582 |   have *: "open (-(\<int>\<^sub>\<le>\<^sub>0 - {z}))"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1583 | by (intro open_Compl discrete_imp_closed[of 1]) (auto elim!: nonpos_Ints_cases simp: dist_of_int) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1584 |   have "Polygamma n holomorphic_on -(\<int>\<^sub>\<le>\<^sub>0 - {z}) - {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1585 | by (intro holomorphic_intros) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1586 | thus "isolated_singularity_at (Polygamma n) z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1587 | by (rule isolated_singularity_at_holomorphic) (use z * in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1588 | qed (auto intro!: holomorphic_intros) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1589 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1590 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1591 | theorem argument_principle': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1592 | fixes f::"complex \<Rightarrow> complex" and poles s:: "complex set" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1593 | \<comment> \<open>\<^term>\<open>pz\<close> is the set of non-essential singularities and zeros\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1594 |   defines "pz \<equiv> {w\<in>s. f w = 0 \<or> w \<in> poles}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1595 | assumes "open s" and | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1596 | "connected s" and | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1597 | f_holo:"f holomorphic_on s-poles" and | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1598 | h_holo:"h holomorphic_on s" and | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1599 | "valid_path g" and | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1600 | loop:"pathfinish g = pathstart g" and | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1601 | path_img:"path_image g \<subseteq> s - pz" and | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1602 | homo:"\<forall>z. (z \<notin> s) \<longrightarrow> winding_number g z = 0" and | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1603 | finite:"finite pz" and | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1604 | poles:"\<forall>p\<in>s\<inter>poles. not_essential f p" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1605 | shows "contour_integral g (\<lambda>x. deriv f x * h x / f x) = 2 * pi * \<i> * | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1606 | (\<Sum>p\<in>pz. winding_number g p * h p * zorder f p)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1607 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1608 | define ff where "ff = remove_sings f" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1609 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1610 | have finite':"finite (s \<inter> poles)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1611 | using finite unfolding pz_def by (auto elim:rev_finite_subset) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1612 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1613 | have isolated:"isolated_singularity_at f z" if "z\<in>s" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1614 | proof (rule isolated_singularity_at_holomorphic) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1615 |     show "f holomorphic_on (s-(poles-{z})) - {z}" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1616 | by (metis Diff_empty Diff_insert Diff_insert0 Diff_subset | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1617 | f_holo holomorphic_on_subset insert_Diff) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1618 |     show "open (s - (poles - {z}))" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1619 | by (metis Diff_Diff_Int Int_Diff assms(2) finite' finite_Diff | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1620 | finite_imp_closed inf.idem open_Diff) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1621 |     show "z \<in> s - (poles - {z})" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1622 | using assms(4) that by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1623 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1624 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1625 | have not_ess:"not_essential f w" if "w\<in>s" for w | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1626 | by (metis Diff_Diff_Int Diff_iff Int_Diff Int_absorb assms(2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1627 | f_holo finite' finite_imp_closed not_essential_holomorphic | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1628 | open_Diff poles that) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1629 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1630 | have nzero:"\<forall>\<^sub>F x in at w. f x \<noteq> 0" if "w\<in>s" for w | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1631 | proof (rule ccontr) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1632 | assume "\<not> (\<forall>\<^sub>F x in at w. f x \<noteq> 0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1633 | then have "\<exists>\<^sub>F x in at w. f x = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1634 | unfolding not_eventually by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1635 | moreover have "\<forall>\<^sub>F x in at w. x\<in>s" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1636 | by (simp add: assms(2) eventually_at_in_open' that) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1637 |     ultimately have "\<exists>\<^sub>F x in at w. x\<in>{w\<in>s. f w = 0}" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1638 | apply (elim frequently_rev_mp) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1639 | by (auto elim:eventually_mono) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1640 | from frequently_at_imp_islimpt[OF this] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1641 |     have "w islimpt {w \<in> s. f w = 0}" .
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1642 |     then have "infinite({w \<in> s. f w = 0} \<inter> ball w 1)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1643 | unfolding islimpt_eq_infinite_ball by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1644 |     then have "infinite({w \<in> s. f w = 0})"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1645 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1646 | then have "infinite pz" unfolding pz_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1647 | by (smt (verit) Collect_mono_iff rev_finite_subset) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1648 | then show False using finite by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1649 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1650 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1651 | obtain pts' where pts':"pts' \<subseteq> s \<inter> poles" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1652 | "finite pts'" "ff holomorphic_on s - pts'" "\<forall>x\<in>pts'. is_pole ff x" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1653 | apply (elim get_all_poles_from_remove_sings | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1654 | [of f,folded ff_def,rotated -1]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1655 | subgoal using f_holo by fastforce | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1656 | using \<open>open s\<close> poles finite' by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1657 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1658 |   have pts'_sub_pz:"{w \<in> s. ff w = 0 \<or> w \<in> pts'} \<subseteq> pz"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1659 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1660 | have "w\<in>poles" if "w\<in>s" "w\<in>pts'" for w | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1661 | by (meson in_mono le_infE pts'(1) that(2)) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1662 | moreover have "f w=0" if" w\<in>s" "w\<notin>poles" "ff w=0" for w | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1663 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1664 | have "\<not> is_pole f w" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1665 | by (metis DiffI Diff_Diff_Int Diff_subset assms(2) f_holo | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1666 | finite' finite_imp_closed inf.absorb_iff2 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1667 | not_is_pole_holomorphic open_Diff that(1) that(2)) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1668 | then have "f \<midarrow>w\<rightarrow> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1669 | using remove_sings_eq_0_iff[OF not_ess[OF \<open>w\<in>s\<close>]] \<open>ff w=0\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1670 | unfolding ff_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1671 |       moreover have "f analytic_on {w}" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1672 | using that(1,2) finite' f_holo assms(2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1673 | by (metis Diff_Diff_Int Diff_empty Diff_iff Diff_subset | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1674 | double_diff finite_imp_closed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1675 | holomorphic_on_imp_analytic_at open_Diff) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1676 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1677 | using ff_def remove_sings_at_analytic that(3) by presburger | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1678 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1679 | ultimately show ?thesis unfolding pz_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1680 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1681 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1682 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1683 | have "contour_integral g (\<lambda>x. deriv f x * h x / f x) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1684 | = contour_integral g (\<lambda>x. deriv ff x * h x / ff x)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1685 | proof (rule contour_integral_eq) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1686 | fix x assume "x \<in> path_image g" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1687 |     have "f analytic_on {x}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1688 | proof (rule holomorphic_on_imp_analytic_at[of _ "s-poles"]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1689 | from finite' | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1690 | show "open (s - poles)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1691 | using \<open>open s\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1692 | by (metis Diff_Compl Diff_Diff_Int Diff_eq finite_imp_closed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1693 | open_Diff) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1694 | show "x \<in> s - poles" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1695 | using path_img \<open>x \<in> path_image g\<close> unfolding pz_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1696 | qed (use f_holo in simp) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1697 | then show "deriv f x * h x / f x = deriv ff x * h x / ff x" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1698 | unfolding ff_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1699 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1700 | also have "... = complex_of_real (2 * pi) * \<i> * | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1701 |                       (\<Sum>p\<in>{w \<in> s. ff w = 0 \<or> w \<in> pts'}. 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1702 | winding_number g p * h p * of_int (zorder ff p))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1703 | proof (rule argument_principle[OF \<open>open s\<close> \<open>connected s\<close>, of ff pts' h g]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1704 |     show "path_image g \<subseteq> s - {w \<in> s. ff w = 0 \<or> w \<in> pts'}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1705 | using path_img pts'_sub_pz by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1706 |     show "finite {w \<in> s. ff w = 0 \<or> w \<in> pts'}" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1707 | using pts'_sub_pz finite | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1708 | using rev_finite_subset by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1709 | qed (use pts' assms in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1710 | also have "... = 2 * pi * \<i> * | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1711 | (\<Sum>p\<in>pz. winding_number g p * h p * zorder f p)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1712 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1713 |     have "(\<Sum>p\<in>{w \<in> s. ff w = 0 \<or> w \<in> pts'}.
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1714 | winding_number g p * h p * of_int (zorder ff p)) = | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1715 | (\<Sum>p\<in>pz. winding_number g p * h p * of_int (zorder f p))" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1716 | proof (rule sum.mono_neutral_cong_left) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1717 | have "zorder f w = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1718 | if "w\<in>s" " f w = 0 \<or> w \<in> poles" "ff w \<noteq> 0" " w \<notin> pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1719 | for w | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1720 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1721 | define F where "F=laurent_expansion f w" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1722 | have has_l:"(\<lambda>x. f (w + x)) has_laurent_expansion F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1723 | unfolding F_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1724 | apply (rule not_essential_has_laurent_expansion) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1725 | using isolated not_ess \<open>w\<in>s\<close> by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1726 | from has_laurent_expansion_eventually_nonzero_iff[OF this] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1727 | have "F \<noteq>0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1728 | using nzero \<open>w\<in>s\<close> by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1729 | from tendsto_0_subdegree_iff[OF has_l this] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1730 | have "f \<midarrow>w\<rightarrow> 0 = (0 < fls_subdegree F)" . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1731 | moreover have "\<not> (is_pole f w \<or> f \<midarrow>w\<rightarrow> 0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1732 | using remove_sings_eq_0_iff[OF not_ess[OF \<open>w\<in>s\<close>]] \<open>ff w \<noteq> 0\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1733 | unfolding ff_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1734 | moreover have "is_pole f w = (fls_subdegree F < 0)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1735 | using is_pole_fls_subdegree_iff[OF has_l] . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1736 | ultimately have "fls_subdegree F = 0" by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1737 | then show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1738 | using has_laurent_expansion_zorder[OF has_l \<open>F\<noteq>0\<close>] by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1739 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1740 |       then show "\<forall>i\<in>pz - {w \<in> s. ff w = 0 \<or> w \<in> pts'}.
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1741 | winding_number g i * h i * of_int (zorder f i) = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1742 | unfolding pz_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1743 |       show "\<And>x. x \<in> {w \<in> s. ff w = 0 \<or> w \<in> pts'} \<Longrightarrow>
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1744 | winding_number g x * h x * of_int (zorder ff x) = | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1745 | winding_number g x * h x * of_int (zorder f x)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1746 | using isolated zorder_remove_sings[of f,folded ff_def] by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1747 | qed (use pts'_sub_pz finite in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1748 | then show ?thesis by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1749 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1750 | finally show ?thesis . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1751 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1752 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1753 | lemma meromorphic_on_imp_isolated_singularity: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1754 | assumes "f meromorphic_on D pts" "z \<in> D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1755 | shows "isolated_singularity_at f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1756 | by (meson DiffI assms(1) assms(2) holomorphic_on_imp_analytic_at isolated_singularity_at_analytic | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1757 | meromorphic_imp_open_diff meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1758 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1759 | lemma meromorphic_imp_not_is_pole: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1760 | assumes "f meromorphic_on D pts" "z \<in> D - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1761 | shows "\<not>is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1762 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1763 |   from assms have "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1764 | using meromorphic_on_imp_analytic_at by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1765 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1766 | using analytic_at not_is_pole_holomorphic by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1767 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1768 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1769 | lemma meromorphic_all_poles_iff_empty [simp]: "f meromorphic_on pts pts \<longleftrightarrow> pts = {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1770 | by (auto simp: meromorphic_on_def holomorphic_on_def open_imp_islimpt) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1771 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1772 | lemma meromorphic_imp_nonsingular_point_exists: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1773 |   assumes "f meromorphic_on A pts" "A \<noteq> {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1774 | obtains x where "x \<in> A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1775 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1776 | have "A \<noteq> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1777 | using assms by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1778 | moreover have "pts \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1779 | using assms by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1780 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1781 | using that by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1782 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1783 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1784 | lemma meromorphic_frequently_const_imp_const: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1785 | assumes "f meromorphic_on A pts" "connected A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1786 | assumes "frequently (\<lambda>w. f w = c) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1787 | assumes "z \<in> A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1788 | assumes "w \<in> A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1789 | shows "f w = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1790 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1791 | have "f w - c = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1792 | proof (rule analytic_continuation[where f = "\<lambda>z. f z - c"]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1793 | show "(\<lambda>z. f z - c) holomorphic_on (A - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1794 | by (intro holomorphic_intros meromorphic_imp_holomorphic[OF assms(1)]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1795 | show [intro]: "open (A - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1796 | using assms meromorphic_imp_open_diff by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1797 | show "connected (A - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1798 | using assms meromorphic_imp_connected_diff by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1799 |     show "{z\<in>A-pts. f z = c} \<subseteq> A - pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1800 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1801 | have "eventually (\<lambda>z. z \<in> A - pts) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1802 | using assms by (intro eventually_at_in_open') auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1803 | hence "frequently (\<lambda>z. f z = c \<and> z \<in> A - pts) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1804 | by (intro frequently_eventually_frequently assms) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1805 |     thus "z islimpt {z\<in>A-pts. f z = c}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1806 | by (simp add: islimpt_conv_frequently_at conj_commute) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1807 | qed (use assms in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1808 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1809 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1810 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1811 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1812 | lemma meromorphic_imp_eventually_neq: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1813 | assumes "f meromorphic_on A pts" "connected A" "\<not>f constant_on A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1814 | assumes "z \<in> A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1815 | shows "eventually (\<lambda>z. f z \<noteq> c) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1816 | proof (rule ccontr) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1817 | assume "\<not>eventually (\<lambda>z. f z \<noteq> c) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1818 | hence *: "frequently (\<lambda>z. f z = c) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1819 | by (auto simp: frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1820 | have "\<forall>w\<in>A-pts. f w = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1821 | using meromorphic_frequently_const_imp_const [OF assms(1,2) * assms(4)] by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1822 | hence "f constant_on A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1823 | by (auto simp: constant_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1824 | thus False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1825 | using assms(3) by contradiction | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1826 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1827 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1828 | lemma meromorphic_frequently_const_imp_const': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1829 | assumes "f meromorphic_on A pts" "connected A" "\<forall>w\<in>pts. is_pole f w" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1830 | assumes "frequently (\<lambda>w. f w = c) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1831 | assumes "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1832 | assumes "w \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1833 | shows "f w = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1834 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1835 | have "\<not>is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1836 | using frequently_const_imp_not_is_pole[OF assms(4)] . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1837 | with assms have z: "z \<in> A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1838 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1839 | have *: "f w = c" if "w \<in> A - pts" for w | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1840 | using that meromorphic_frequently_const_imp_const [OF assms(1,2,4) z] by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1841 | have "\<not>is_pole f u" if "u \<in> A" for u | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1842 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1843 | have "is_pole f u \<longleftrightarrow> is_pole (\<lambda>_. c) u" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1844 | proof (rule is_pole_cong) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1845 |       have "eventually (\<lambda>w. w \<in> A - (pts - {u}) - {u}) (at u)"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1846 | by (intro eventually_at_in_open meromorphic_imp_open_diff' [OF assms(1)]) (use that in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1847 | thus "eventually (\<lambda>w. f w = c) (at u)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1848 | by eventually_elim (use * in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1849 | qed auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1850 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1851 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1852 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1853 | moreover have "pts \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1854 | using assms(1) by (simp add: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1855 |   ultimately have "pts = {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1856 | using assms(3) by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1857 | with * and \<open>w \<in> A\<close> show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1858 | by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1859 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1860 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1861 | lemma meromorphic_imp_eventually_neq': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1862 | assumes "f meromorphic_on A pts" "connected A" "\<forall>w\<in>pts. is_pole f w" "\<not>f constant_on A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1863 | assumes "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1864 | shows "eventually (\<lambda>z. f z \<noteq> c) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1865 | proof (rule ccontr) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1866 | assume "\<not>eventually (\<lambda>z. f z \<noteq> c) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1867 | hence *: "frequently (\<lambda>z. f z = c) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1868 | by (auto simp: frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1869 | have "\<forall>w\<in>A. f w = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1870 | using meromorphic_frequently_const_imp_const' [OF assms(1,2,3) * assms(5)] by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1871 | hence "f constant_on A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1872 | by (auto simp: constant_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1873 | thus False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1874 | using assms(4) by contradiction | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1875 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1876 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1877 | lemma zorder_eq_0_iff_meromorphic: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1878 | assumes "f meromorphic_on A pts" "\<forall>z\<in>pts. is_pole f z" "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1879 | assumes "eventually (\<lambda>x. f x \<noteq> 0) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1880 | shows "zorder f z = 0 \<longleftrightarrow> \<not>is_pole f z \<and> f z \<noteq> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1881 | proof (cases "z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1882 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1883 | from assms obtain F where F: "(\<lambda>x. f (z + x)) has_laurent_expansion F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1884 | by (metis True meromorphic_on_def not_essential_has_laurent_expansion) (* TODO: better lemmas *) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1885 | from F and assms(4) have [simp]: "F \<noteq> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1886 | using has_laurent_expansion_eventually_nonzero_iff by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1887 | show ?thesis using True assms(2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1888 | using is_pole_fls_subdegree_iff [OF F] has_laurent_expansion_zorder [OF F] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1889 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1890 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1891 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1892 |   have ana: "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1893 | using meromorphic_on_imp_analytic_at False assms by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1894 | hence "\<not>is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1895 | using analytic_at not_is_pole_holomorphic by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1896 | moreover have "frequently (\<lambda>w. f w \<noteq> 0) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1897 | using assms(4) by (intro eventually_frequently) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1898 | ultimately show ?thesis using zorder_eq_0_iff[OF ana] False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1899 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1900 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1901 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1902 | lemma zorder_pos_iff_meromorphic: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1903 | assumes "f meromorphic_on A pts" "\<forall>z\<in>pts. is_pole f z" "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1904 | assumes "eventually (\<lambda>x. f x \<noteq> 0) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1905 | shows "zorder f z > 0 \<longleftrightarrow> \<not>is_pole f z \<and> f z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1906 | proof (cases "z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1907 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1908 | from assms obtain F where F: "(\<lambda>x. f (z + x)) has_laurent_expansion F" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1909 | by (metis True meromorphic_on_def not_essential_has_laurent_expansion) (* TODO: better lemmas *) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1910 | from F and assms(4) have [simp]: "F \<noteq> 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1911 | using has_laurent_expansion_eventually_nonzero_iff by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1912 | show ?thesis using True assms(2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1913 | using is_pole_fls_subdegree_iff [OF F] has_laurent_expansion_zorder [OF F] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1914 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1915 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1916 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1917 |   have ana: "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1918 | using meromorphic_on_imp_analytic_at False assms by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1919 | hence "\<not>is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1920 | using analytic_at not_is_pole_holomorphic by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1921 | moreover have "frequently (\<lambda>w. f w \<noteq> 0) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1922 | using assms(4) by (intro eventually_frequently) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1923 | ultimately show ?thesis using zorder_pos_iff'[OF ana] False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1924 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1925 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1926 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1927 | lemma zorder_neg_iff_meromorphic: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1928 | assumes "f meromorphic_on A pts" "\<forall>z\<in>pts. is_pole f z" "z \<in> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1929 | assumes "eventually (\<lambda>x. f x \<noteq> 0) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1930 | shows "zorder f z < 0 \<longleftrightarrow> is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1931 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1932 | have "frequently (\<lambda>x. f x \<noteq> 0) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1933 | using assms by (intro eventually_frequently) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1934 | moreover from assms have "isolated_singularity_at f z" "not_essential f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1935 | using meromorphic_on_imp_isolated_singularity meromorphic_on_imp_not_essential by blast+ | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1936 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1937 | using isolated_pole_imp_neg_zorder neg_zorder_imp_is_pole by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1938 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1939 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1940 | lemma meromorphic_on_imp_discrete: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1941 | assumes mero:"f meromorphic_on S pts" and "connected S" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1942 | and nconst:"\<not> (\<forall>w\<in>S - pts. f w = c)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1943 |   shows "discrete {x\<in>S. f x=c}" 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1944 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1945 | define g where "g=(\<lambda>x. f x - c)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1946 | have "\<forall>\<^sub>F w in at z. g w \<noteq> 0" if "z \<in> S" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1947 | proof (rule nconst_imp_nzero_neighbour'[of g S pts z]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1948 | show "g meromorphic_on S pts" using mero unfolding g_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1949 | by (auto intro:meromorphic_intros) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1950 | show "\<not> (\<forall>w\<in>S - pts. g w = 0)" using nconst unfolding g_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1951 | qed fact+ | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1952 | then show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1953 | unfolding discrete_altdef g_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1954 | using eventually_mono by fastforce | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1955 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1956 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1957 | lemma meromorphic_isolated_in: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1958 | assumes merf: "f meromorphic_on D pts" "p\<in>pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1959 | shows "p isolated_in pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1960 | by (meson assms isolated_in_islimpt_iff meromorphic_on_def subsetD) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1961 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1962 | lemma remove_sings_constant_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1963 | assumes merf: "f meromorphic_on D pts" and "connected D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1964 | and const:"f constant_on (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1965 | shows "(remove_sings f) constant_on D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1966 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1967 | have remove_sings_const: "remove_sings f constant_on D - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1968 | using const | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1969 | by (metis constant_onE merf meromorphic_on_imp_analytic_at remove_sings_at_analytic) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1970 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1971 |   have ?thesis if "D = {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1972 | using that unfolding constant_on_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1973 |   moreover have ?thesis if "D\<noteq>{}" "{x\<in>pts. is_pole f x} = {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1974 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1975 | obtain \<xi> where "\<xi> \<in> (D - pts)" "\<xi> islimpt (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1976 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1977 | have "open (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1978 | using meromorphic_imp_open_diff[OF merf] . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1979 |       moreover have "(D - pts) \<noteq> {}" using \<open>D\<noteq>{}\<close>
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1980 | by (metis Diff_empty closure_empty merf | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1981 | meromorphic_pts_closure subset_empty) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1982 | ultimately show ?thesis using open_imp_islimpt that by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1983 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1984 | moreover have "remove_sings f holomorphic_on D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1985 | using remove_sings_holomorphic_on[OF merf] that by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1986 | moreover note remove_sings_const | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1987 | moreover have "open D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1988 | using assms(1) meromorphic_on_def by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1989 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1990 | using Conformal_Mappings.analytic_continuation' | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1991 | [of "remove_sings f" D "D-pts" \<xi>] \<open>connected D\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1992 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1993 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1994 |   moreover have ?thesis if "D\<noteq>{}" "{x\<in>pts. is_pole f x} \<noteq> {}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1995 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1996 |     define PP where "PP={x\<in>D. is_pole f x}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1997 | have "remove_sings f meromorphic_on D PP" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1998 | using merf unfolding PP_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 1999 | apply (elim remove_sings_meromorphic_on) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2000 | subgoal using assms(1) meromorphic_on_def by force | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2001 | subgoal using meromorphic_pole_subset merf by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2002 | done | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2003 | moreover have "remove_sings f constant_on D - PP" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2004 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2005 | obtain \<xi> where "\<xi> \<in> f ` (D - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2006 |         by (metis Diff_empty Diff_eq_empty_iff \<open>D \<noteq> {}\<close> assms(1) 
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2007 | closure_empty ex_in_conv imageI meromorphic_pts_closure) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2008 | have \<xi>:"\<forall>x\<in>D - pts. f x = \<xi>" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2009 | by (metis \<open>\<xi> \<in> f ` (D - pts)\<close> assms(3) constant_on_def image_iff) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2010 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2011 | have "remove_sings f x = \<xi>" if "x\<in>D - PP" for x | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2012 | proof (cases "x\<in>pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2013 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2014 | then have"x isolated_in pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2015 | using meromorphic_isolated_in[OF merf] by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2016 |         then obtain T0 where T0:"open T0" "T0 \<inter> pts = {x}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2017 | unfolding isolated_in_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2018 | obtain T1 where T1:"open T1" "x\<in>T1" "T1 \<subseteq> D" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2019 | using merf unfolding meromorphic_on_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2020 | using True by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2021 | define T2 where "T2 = T1 \<inter> T0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2022 |         have "open T2" "x\<in>T2" "T2 - {x} \<subseteq> D - pts"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2023 | using T0 T1 unfolding T2_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2024 | then have "\<forall>w\<in>T2. w\<noteq>x \<longrightarrow> f w =\<xi>" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2025 | using \<xi> by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2026 | then have "\<forall>\<^sub>F x in at x. f x = \<xi>" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2027 | unfolding eventually_at_topological | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2028 | using \<open>open T2\<close> \<open>x\<in>T2\<close> by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2029 | then have "f \<midarrow>x\<rightarrow> \<xi>" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2030 | using tendsto_eventually by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2031 | then show ?thesis by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2032 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2033 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2034 | then show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2035 | using \<open>\<forall>x\<in>D - pts. f x = \<xi>\<close> assms(1) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2036 | meromorphic_on_imp_analytic_at that by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2037 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2038 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2039 | then show ?thesis unfolding constant_on_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2040 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2041 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2042 | moreover have "is_pole (remove_sings f) x" if "x\<in>PP" for x | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2043 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2044 | have "isolated_singularity_at f x" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2045 | by (metis (mono_tags, lifting) DiffI PP_def assms(1) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2046 | isolated_singularity_at_analytic mem_Collect_eq | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2047 | meromorphic_on_def meromorphic_on_imp_analytic_at that) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2048 | then show ?thesis using that unfolding PP_def by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2049 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2050 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2051 | using meromorphic_imp_constant_on | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2052 | [of "remove_sings f" D PP] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2053 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2054 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2055 | ultimately show ?thesis by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2056 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2057 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2058 | lemma meromorphic_eq_meromorphic_extend: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2059 | assumes "f meromorphic_on A pts1" "g meromorphic_on A pts1" "\<not>z islimpt pts2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2060 | assumes "\<And>z. z \<in> A - pts2 \<Longrightarrow> f z = g z" "pts1 \<subseteq> pts2" "z \<in> A - pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2061 | shows "f z = g z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2062 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2063 |   have "g analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2064 | using assms by (intro meromorphic_on_imp_analytic_at[OF assms(2)]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2065 | hence "g \<midarrow>z\<rightarrow> g z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2066 | using analytic_at_imp_isCont isContD by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2067 | also have "?this \<longleftrightarrow> f \<midarrow>z\<rightarrow> g z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2068 | proof (intro filterlim_cong) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2069 | have "eventually (\<lambda>w. w \<notin> pts2) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2070 | using assms by (auto simp: islimpt_conv_frequently_at frequently_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2071 | moreover have "eventually (\<lambda>w. w \<in> A) (at z)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2072 | using assms by (intro eventually_at_in_open') (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2073 | ultimately show "\<forall>\<^sub>F x in at z. g x = f x" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2074 | by eventually_elim (use assms in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2075 | qed auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2076 | finally have "f \<midarrow>z\<rightarrow> g z" . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2077 |   moreover have "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2078 | using assms by (intro meromorphic_on_imp_analytic_at[OF assms(1)]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2079 | hence "f \<midarrow>z\<rightarrow> f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2080 | using analytic_at_imp_isCont isContD by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2081 | ultimately show ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2082 | using tendsto_unique by force | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2083 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2084 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2085 | lemma meromorphic_constant_on_extend: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2086 | assumes "f constant_on A - pts1" "f meromorphic_on A pts1" "f meromorphic_on A pts2" "pts2 \<subseteq> pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2087 | shows "f constant_on A - pts2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2088 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2089 | from assms(1) obtain c where c: "\<And>z. z \<in> A - pts1 \<Longrightarrow> f z = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2090 | unfolding constant_on_def by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2091 | have "f z = c" if "z \<in> A - pts2" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2092 | using assms(3) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2093 | proof (rule meromorphic_eq_meromorphic_extend[where z = z]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2094 | show "(\<lambda>a. c) meromorphic_on A pts2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2095 | by (intro meromorphic_on_const) (use assms in \<open>auto simp: meromorphic_on_def\<close>) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2096 | show "\<not>z islimpt pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2097 | using that assms by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2098 | qed (use assms c that in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2099 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2100 | by (auto simp: constant_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2101 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2102 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2103 | lemma meromorphic_remove_sings_constant_on_imp_constant_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2104 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2105 | assumes "remove_sings f constant_on A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2106 | shows "f constant_on A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2107 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2108 | from assms(2) obtain c where c: "\<And>z. z \<in> A \<Longrightarrow> remove_sings f z = c" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2109 | by (auto simp: constant_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2110 | have "f z = c" if "z \<in> A - pts" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2111 | using meromorphic_on_imp_analytic_at[OF assms(1) that] c[of z] that | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2112 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2113 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2114 | by (auto simp: constant_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2115 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2116 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2117 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2118 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2119 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2120 | definition singularities_on :: "complex set \<Rightarrow> (complex \<Rightarrow> complex) \<Rightarrow> complex set" where | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2121 | "singularities_on A f = | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2122 |      {z\<in>A. isolated_singularity_at f z \<and> not_essential f z \<and> \<not>f analytic_on {z}}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2123 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2124 | lemma singularities_on_subset: "singularities_on A f \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2125 | by (auto simp: singularities_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2126 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2127 | lemma pole_in_singularities_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2128 | assumes "f meromorphic_on A pts" "z \<in> A" "is_pole f z" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2129 | shows "z \<in> singularities_on A f" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2130 | unfolding singularities_on_def not_essential_def using assms | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2131 | using analytic_at_imp_no_pole meromorphic_on_imp_isolated_singularity by force | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2132 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2133 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2134 | lemma meromorphic_on_subset_pts: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2135 | assumes "f meromorphic_on A pts" "pts' \<subseteq> pts" "f analytic_on pts - pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2136 | shows "f meromorphic_on A pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2137 | proof | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2138 | show "open A" "pts' \<subseteq> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2139 | using assms by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2140 | show "isolated_singularity_at f z" "not_essential f z" if "z \<in> pts'" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2141 | using assms that by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2142 | show "\<not>z islimpt pts'" if "z \<in> A" for z | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2143 | using assms that islimpt_subset unfolding meromorphic_on_def by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2144 | have "f analytic_on A - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2145 | using assms(1) meromorphic_imp_analytic by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2146 | with assms have "f analytic_on (A - pts) \<union> (pts - pts')" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2147 | by (subst analytic_on_Un) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2148 | also have "(A - pts) \<union> (pts - pts') = A - pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2149 | using assms by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2150 | finally show "f holomorphic_on A - pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2151 | using analytic_imp_holomorphic by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2152 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2153 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2154 | lemma meromorphic_on_imp_superset_singularities_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2155 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2156 | shows "singularities_on A f \<subseteq> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2157 | proof | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2158 | fix z assume "z \<in> singularities_on A f" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2159 |   hence "z \<in> A" "\<not>f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2160 | by (auto simp: singularities_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2161 | with assms show "z \<in> pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2162 | by (meson DiffI meromorphic_on_imp_analytic_at) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2163 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2164 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2165 | lemma meromorphic_on_singularities_on: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2166 | assumes "f meromorphic_on A pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2167 | shows "f meromorphic_on A (singularities_on A f)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2168 | using assms meromorphic_on_imp_superset_singularities_on[OF assms] | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2169 | proof (rule meromorphic_on_subset_pts) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2170 |   have "f analytic_on {z}" if "z \<in> pts - singularities_on A f" for z
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2171 | using that assms by (auto simp: singularities_on_def meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2172 | thus "f analytic_on pts - singularities_on A f" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2173 | using analytic_on_analytic_at by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2174 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2175 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2176 | theorem Residue_theorem_inside: | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2177 | assumes f: "f meromorphic_on s pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2178 | "simply_connected s" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2179 | assumes g: "valid_path g" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2180 | "pathfinish g = pathstart g" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2181 | "path_image g \<subseteq> s - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2182 | defines "pts1 \<equiv> pts \<inter> inside (path_image g)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2183 | shows "finite pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2184 | and "contour_integral g f = 2 * pi * \<i> * (\<Sum>p\<in>pts1. winding_number g p * residue f p)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2185 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2186 | note [dest] = valid_path_imp_path | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2187 | have cl_g [intro]: "closed (path_image g)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2188 | using g by (auto intro!: closed_path_image) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2189 | have "open s" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2190 | using f(1) by (auto simp: meromorphic_on_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2191 | define pts2 where "pts2 = pts - pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2192 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2193 | define A where "A = path_image g \<union> inside (path_image g)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2194 | have "closed A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2195 | unfolding A_def using g by (intro closed_path_image_Un_inside) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2196 | moreover have "bounded A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2197 | unfolding A_def using g by (auto intro!: bounded_path_image bounded_inside) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2198 | ultimately have 1: "compact A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2199 | using compact_eq_bounded_closed by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2200 | have 2: "open (s - pts2)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2201 | using f by (auto intro!: meromorphic_imp_open_diff' [OF f(1)] simp: pts2_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2202 | have 3: "A \<subseteq> s - pts2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2203 | unfolding A_def pts2_def pts1_def | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2204 | using f(2) g(3) 2 subset_simply_connected_imp_inside_subset[of s "path_image g"] \<open>open s\<close> | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2205 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2206 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2207 | obtain \<epsilon> where \<epsilon>: "\<epsilon> > 0" "(\<Union>x\<in>A. ball x \<epsilon>) \<subseteq> s - pts2" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2208 | using compact_subset_open_imp_ball_epsilon_subset[OF 1 2 3] by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2209 | define B where "B = (\<Union>x\<in>A. ball x \<epsilon>)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2210 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2211 | have "finite (A \<inter> pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2212 | using 1 3 by (intro meromorphic_compact_finite_pts[OF f(1)]) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2213 | also have "A \<inter> pts = pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2214 | unfolding pts1_def using g by (auto simp: A_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2215 | finally show fin: "finite pts1" . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2216 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2217 | show "contour_integral g f = 2 * pi * \<i> * (\<Sum>p\<in>pts1. winding_number g p * residue f p)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2218 | proof (rule Residue_theorem) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2219 | show "open B" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2220 | by (auto simp: B_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2221 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2222 | have "connected A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2223 | unfolding A_def using g | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2224 | by (intro connected_with_inside closed_path_image connected_path_image) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2225 | hence "connected (A \<union> B)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2226 | unfolding B_def using g \<open>\<epsilon> > 0\<close> f(2) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2227 | by (intro connected_Un_UN connected_path_image valid_path_imp_path) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2228 | (auto simp: simply_connected_imp_connected) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2229 | also have "A \<union> B = B" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2230 | using \<epsilon>(1) by (auto simp: B_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2231 | finally show "connected B" . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2232 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2233 | have "f holomorphic_on (s - pts)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2234 | by (intro meromorphic_imp_holomorphic f) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2235 | moreover have "B - pts1 \<subseteq> s - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2236 | using \<epsilon> unfolding B_def by (auto simp: pts1_def pts2_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2237 | ultimately show "f holomorphic_on (B - pts1)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2238 | by (rule holomorphic_on_subset) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2239 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2240 | have "path_image g \<subseteq> A - pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2241 | using g unfolding pts1_def by (auto simp: A_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2242 | also have "\<dots> \<subseteq> B - pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2243 | unfolding B_def using \<epsilon>(1) by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2244 | finally show "path_image g \<subseteq> B - pts1" . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2245 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2246 | show "\<forall>z. z \<notin> B \<longrightarrow> winding_number g z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2247 | proof safe | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2248 | fix z assume z: "z \<notin> B" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2249 | hence "z \<notin> A" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2250 | using \<epsilon>(1) by (auto simp: B_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2251 | hence "z \<in> outside (path_image g)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2252 | unfolding A_def by (simp add: union_with_inside) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2253 | thus "winding_number g z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2254 | using g by (intro winding_number_zero_in_outside) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2255 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2256 | qed (use g fin in auto) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2257 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2258 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2259 | theorem Residue_theorem': | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2260 | assumes f: "f meromorphic_on s pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2261 | "simply_connected s" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2262 | assumes g: "valid_path g" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2263 | "pathfinish g = pathstart g" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2264 | "path_image g \<subseteq> s - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2265 | assumes pts': "finite pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2266 | "pts' \<subseteq> s" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2267 | "\<And>z. z \<in> pts - pts' \<Longrightarrow> winding_number g z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2268 | shows "contour_integral g f = 2 * pi * \<i> * (\<Sum>p\<in>pts'. winding_number g p * residue f p)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2269 | proof - | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2270 | note [dest] = valid_path_imp_path | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2271 | define pts1 where "pts1 = pts \<inter> inside (path_image g)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2272 | |
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2273 | have "contour_integral g f = 2 * pi * \<i> * (\<Sum>p\<in>pts1. winding_number g p * residue f p)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2274 | unfolding pts1_def by (intro Residue_theorem_inside[OF f g]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2275 | also have "(\<Sum>p\<in>pts1. winding_number g p * residue f p) = | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2276 | (\<Sum>p\<in>pts'. winding_number g p * residue f p)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2277 | proof (intro sum.mono_neutral_cong refl) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2278 | show "finite pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2279 | unfolding pts1_def by (intro Residue_theorem_inside[OF f g]) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2280 | show "finite pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2281 | by fact | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2282 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2283 | fix z assume z: "z \<in> pts' - pts1" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2284 | show "winding_number g z * residue f z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2285 | proof (cases "z \<in> pts") | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2286 | case True | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2287 | with z have "z \<notin> path_image g \<union> inside (path_image g)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2288 | using g(3) by (auto simp: pts1_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2289 | hence "z \<in> outside (path_image g)" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2290 | by (simp add: union_with_inside) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2291 | hence "winding_number g z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2292 | using g by (intro winding_number_zero_in_outside) auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2293 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2294 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2295 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2296 | case False | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2297 | with z pts' have "z \<in> s - pts" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2298 | by auto | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2299 |       with f(1) have "f analytic_on {z}"
 | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2300 | by (intro meromorphic_on_imp_analytic_at) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2301 | hence "residue f z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2302 | using analytic_at residue_holo by blast | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2303 | thus ?thesis | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2304 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2305 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2306 | next | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2307 | fix z assume z: "z \<in> pts1 - pts'" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2308 | hence "winding_number g z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2309 | using pts' by (auto simp: pts1_def) | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2310 | thus "winding_number g z * residue f z = 0" | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2311 | by simp | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2312 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2313 | finally show ?thesis . | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2314 | qed | 
| 
c6b50597abbc
More of Eberl's contributions: memomorphic functions
 paulson <lp15@cam.ac.uk> parents: diff
changeset | 2315 | |
| 78698 | 2316 | end |